Skip to content
GitLab
Menu
Projects
Groups
Snippets
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Nicola Botta
papers
Commits
e6c7245c
Commit
e6c7245c
authored
Feb 23, 2021
by
Nuria Brede
Browse files
Initial.
parent
8e0120fa
Changes
12
Expand all
Hide whitespace changes
Inline
Side-by-side
2021.On the Correctness of Monadic Backward Induction/code/Equality.lidr
0 → 100644
View file @
e6c7245c
> module Equality
> %default total
> %auto_implicits off
> %access public export
> ExtEq : {A, B : Type} -> (f, g : A -> B) -> Type
> ExtEq {A} f g = (a : A) -> f a = g a
2021.On the Correctness of Monadic Backward Induction/code/Framework.lidr
0 → 100644
View file @
e6c7245c
This diff is collapsed.
Click to expand it.
2021.On the Correctness of Monadic Backward Induction/code/GoodBadCounterExamples.lidr
0 → 100644
View file @
e6c7245c
> module GoodBadCounterExamples
> import Data.Fin
> import Equality
> import Framework
> import List
> %default total
> %auto_implicits off
> %access public export
This file contains some of the counter-examples to
|valMeasTotalReward| of [1], section 5.2.
Two "states of the world":
> data StateOfTheWorld = Good | Bad
> Framework.X _ = StateOfTheWorld
> finiteState : {t : Nat} -> Framework.X t -> Fin 2
> finiteState Good = FZ
> finiteState Bad = FS FZ
Two "emission control options":
> data EmissionLevel = Low | High
> Framework.Y _ _ = EmissionLevel
> finiteControl : {t : Nat} -> {x : Framework.X t} -> Framework.Y t x -> Fin 2
> finiteControl Low = FZ
> finiteControl High = FS FZ
< Framework.M = List
For counter-examples with |Val = Nat|:
> Framework.Val = Nat
> Framework.zero = 0
For counter-example with |Val = Double|:
> -- Framework.Val = Double
> -- Framework.zero = 0.0
> Framework.(<+>) = (+)
with |Val = Nat|:
> Framework.meas = foldr (\x, v => (x + 1 `maximum` v)) 0
> -- Framework.meas = foldr (*) 1.0
> -- Framework.meas = foldr (*) 1.0
with |Val = Double|:
> avg : List Double -> Double
> avg [] = 0.0
> avg xs = sum xs / cast (length xs)
> -- Framework.meas = avg
Transition function:
> Framework.next _t Good Low = [Good]
> Framework.next _t Bad High = [Bad]
> Framework.next _t _x _y = [Good, Bad]
Reward/ cost function:
> Framework.reward _t _x Low Good = 3
> Framework.reward _t _x High Good = 2
> Framework.reward _t _x Low Bad = 1
> Framework.reward _t _x High Bad = 0
Constant policies:
> constH : (t : Nat) -> Framework.Policy t
> constH _ = const High
> constL : (t : Nat) -> Framework.Policy t
> constL _ = const Low
Other simple policies:
> react : (t : Nat) -> Framework.Policy t
> react _ Good = High
> react _ Bad = Low
> maintain : (t : Nat) -> Framework.Policy t
> maintain _ Good = Low
> maintain _ Bad = High
Example policy sequences:
> ps : Framework.PolicySeq 0 2
> ps = (constL 0 :: (constH 1 :: Nil))
> ps' : Framework.PolicySeq 0 2
> ps' = (constH 0 :: (constL 1 :: Nil))
> ps1 : Framework.PolicySeq 0 3
> ps1 = constH 0 :: (constL 1 :: (constH 2 :: Nil))
> ps2 : Framework.PolicySeq 0 3
> ps2 = constH 0 :: (constH 1 :: (react 2 :: Nil))
> ps3 : Framework.PolicySeq 0 5
> ps3 = constH 0 :: (maintain 1 :: (react 2 :: (constH 3 :: (react 4 :: Nil))))
Some test cases:
----------------
Initial state:
> x0 : Framework.X 0
> x0 = Good
> -- x0 = Bad
> test : Type
> test = val ps x0 = val' ps x0
> test1 : Type
> test1 = val ps1 x0 = val' ps1 x0
> test2 : Type
> test2 = val ps2 x0 = val' ps2 x0
> test3 : Type
> test3 = val ps3 x0 = val' ps3 x0
2021.On the Correctness of Monadic Backward Induction/code/GoodBadExample.lidr
0 → 100644
View file @
e6c7245c
> module GoodBadExample
> import Data.Fin
> import Equality
> import Framework
> import List
> %default total
> %auto_implicits off
> %access public export
The following example corresponds to the toy "climate" SDP described in section 2.1 of [1].
Two "states of the world":
> data StateOfTheWorld = Good | Bad
> Framework.X _ = StateOfTheWorld
> finiteState : {t : Nat} -> Framework.X t -> Fin 2
> finiteState Good = FZ
> finiteState Bad = FS FZ
Two "emission control options":
> data EmissionLevel = Low | High
> Framework.Y _ _ = EmissionLevel
> finiteControl : {t : Nat} -> {x : Framework.X t} -> Framework.Y t x -> Fin 2
> finiteControl Low = FZ
> finiteControl High = FS FZ
< Framework.M = List
> Framework.Val = Nat
> Framework.zero = 0
> Framework.(<+>) = (+)
> maxList : List Nat -> Nat
> maxList [] = 0
> maxList (x :: xs) = x `maximum` (maxList xs)
> minList : List Nat -> Nat
> minList [] = 0
> minList (x :: []) = x
> minList (x :: xs) = x `minimum` (minList xs)
Comment/uncomment to try the two measures:
> Framework.meas = minList
> -- Framework.meas = maxList
Transition function:
> Framework.next _t Good Low = [Good]
> Framework.next _t Bad High = [Bad]
> Framework.next _t _x _y = [Good, Bad]
Reward/ cost function:
> Framework.reward _t _x Low Good = 3
> Framework.reward _t _x High Good = 2
> Framework.reward _t _x Low Bad = 1
> Framework.reward _t _x High Bad = 0
Constant policies:
> constH : (t : Nat) -> Framework.Policy t
> constH _ = const High
> constL : (t : Nat) -> Framework.Policy t
> constL _ = const Low
Other simple policies:
> react : (t : Nat) -> Framework.Policy t
> react _ Good = High
> react _ Bad = Low
> maintain : (t : Nat) -> Framework.Policy t
> maintain _ Good = Low
> maintain _ Bad = High
Example policy sequences:
> ps : Framework.PolicySeq 0 2
> ps = (constL 0 :: (constH 1 :: Nil))
> ps' : Framework.PolicySeq 0 2
> ps' = (constH 0 :: (constL 1 :: Nil))
> ps1 : Framework.PolicySeq 0 3
> ps1 = constH 0 :: (constL 1 :: (constH 2 :: Nil))
> ps2 : Framework.PolicySeq 0 3
> ps2 = constH 0 :: (constH 1 :: (react 2 :: Nil))
> ps3 : Framework.PolicySeq 0 5
> ps3 = constH 0 :: (maintain 1 :: (react 2 :: (constH 3 :: (react 4 :: Nil))))
Some test cases:
----------------
Initial state:
> x0 : Framework.X 0
> x0 = Good
> -- x0 = Bad
> test : Type
> test = val ps x0 = val' ps x0
> test1 : Type
> test1 = val ps1 x0 = val' ps1 x0
> test2 : Type
> test2 = val ps2 x0 = val' ps2 x0
> test3 : Type
> test3 = val ps3 x0 = val' ps3 x0
2021.On the Correctness of Monadic Backward Induction/code/Id.lidr
0 → 100644
View file @
e6c7245c
> module Id
> import Syntax.PreorderReasoning
> import Equality
> import Framework
> %default total
> %auto_implicits off
> %access public export
We consider an instance of the framework for deterministic
sequential decision problems (thus with degenerated uncertainty),
using the |Id| monad.
All the definitions below are evidently trivial in this case.
> Framework.M T = T
> Framework.map = id
> Framework.pure = id
> Framework.join = id
> Framework.(>>=) ma f = f ma
< mapPresExtEq : {A, B : Type} -> (f, g : A -> B) ->
< f `ExtEq` g -> map f `ExtEq` map g
> Framework.mapPresExtEq {A} {B} f g ee a = ee a
< NotEmpty : {A : Type} -> Id A -> Type
> Framework.NotEmpty a = Unit
< pureNotEmpty : {A : Type} -> (a : A) -> NotEmpty (pure a)
> Framework.pureNotEmpty a = ()
< mapPresNotEmpty : {A, B : Type} -> (f : A -> B) -> (ma : M A) ->
< NotEmpty ma -> NotEmpty (map f ma)
> Framework.mapPresNotEmpty f a _ = ()
< bindPresNotEmpty : {A, B : Type} -> (f : A -> M B) -> (ma : M A) ->
< NotEmpty ma -> -> ((a : A) -> NotEmpty (f a)) ->
< NotEmpty (ma >>= f)
> Framework.bindPresNotEmpty f a ne _ = ()
We omit the proofs of the standard monad laws for |Id|.
2021.On the Correctness of Monadic Backward Induction/code/List.lidr
0 → 100644
View file @
e6c7245c
> module List
> import Syntax.PreorderReasoning
> import Equality
> import Framework
> %default total
> %auto_implicits off
> %access public export
We consider an instance of the framework
for non-deterministic sequential decision problems, using the |List| monad:
> Framework.M = List
< Framework.map f [] = []
< Framework.map f (a :: as) = f a :: map f as
> Framework.map f = foldr (\a => ((f a) ::)) []
> Framework.pure a = [a]
> Framework.join = concat
> Framework.(>>=) ma f = (join . map f) ma
First we need to make sure that |map| preserves extensional equality:
< mapPresExtEq : {A, B : Type} -> (f, g : A -> B) ->
< f `ExtEq` g -> map f `ExtEq` map g
> Framework.mapPresExtEq {A} {B} f g ee [] = Refl
> Framework.mapPresExtEq {A} {B} f g ee (a :: as) = rewrite ee a in
> rewrite mapPresExtEq f g ee as in
> Refl
Then we need to define when a list is non-empty:
< NotEmpty : {A : Type} -> List A -> Type
> Framework.NotEmpty Nil = Void
> Framework.NotEmpty (a :: as) = Unit
We have to show that |pure| fulfills the non-emptiness requirement:
< pureNotEmpty : {A : Type} -> (a : A) -> NotEmpty (pure a)
> Framework.pureNotEmpty a = ()
And that |map| and |(>>=)| preserve non-emptiness:
< mapPresNotEmpty : {A, B : Type} -> (f : A -> B) -> (ma : M A) ->
< NotEmpty ma -> NotEmpty (map f ma)
> Framework.mapPresNotEmpty f [] ne = void ne
> Framework.mapPresNotEmpty f (a :: as) _ = ()
To show |bindPresNotEmptySpec| we first prove a lemma about preservation of
non-emptiness by |++|:
> appendPresNotEmptyLeft : {A : Type} -> (as, as' : List A) -> NotEmpty as ->
> NotEmpty (as ++ as')
> appendPresNotEmptyLeft [] _ ne = void ne
> appendPresNotEmptyLeft as [] ne = rewrite appendNilRightNeutral as in ne
> appendPresNotEmptyLeft (a :: as) _ _ = ()
Then
< bindPresNotEmpty : {A, B : Type} -> (f : A -> M B) -> (ma : M A) ->
< NotEmpty ma -> -> ((a : A) -> NotEmpty (f a)) ->
< NotEmpty (ma >>= f)
> Framework.bindPresNotEmpty f [] ne _ = void ne
> Framework.bindPresNotEmpty f (a :: as) _ nefa =
> appendPresNotEmptyLeft (f a) (join (map f as)) (nefa a)
We omit the proofs of the standard monad laws for |List|.
2021.On the Correctness of Monadic Backward Induction/code/ListMeasures.lidr
0 → 100644
View file @
e6c7245c
> module ListMeasures
> import Syntax.PreorderReasoning
> import Equality
> import Framework
> import List
> %default total
> %auto_implicits off
> %access public export
--------------------------------------------------------------------------
In this file we show that for monoids |(Val, odot, neutr)| and measures
that are monoid homomorphisms |meas = foldr odot neutr|
the three conditions hold, if |<+>| distributes over |odot| on the left.
(Otherwise just |measPureSpec| and |measJoinSpec| hold.)
We also show that |measMonSpec| holds if |odot| is monotone with
respect to |<=|.
--------------------------------------------------------------------------
Define
> odot : Val -> Val -> Val
> neutr : Val
> odotList : List Val -> Val
which can be written as |List|-fold
> odotList = foldr odot neutr
and consider the following (partial) instance of the framework:
< Framework.M = List -- imported from |List.lidr|
> Framework.meas = odotList
We require |(Val, odot, neutr)| to be a monoid.
> odotNeutrRight : (l : Val) -> l `odot` neutr = l
> odotNeutrLeft : (r : Val) -> neutr `odot` r = r
>
> odotAssociative : (l, v, r : Val) -> l `odot` (v `odot` r) = (l `odot` v) `odot` r
> oplusOdotDistrLeft : (n, l, r : Val) -> n <+> (l `odot` r) = n <+> l `odot` n <+> r
Then |odotList| is a monoid homomorphism and thus preserves the binary operation.
> odotListAppend : (xs, xs' : List Val) ->
> odotList (xs ++ xs') = ((odotList xs) `odot` (odotList xs'))
> odotListAppend [] xs' = rewrite odotNeutrLeft (odotList xs') in Refl
> odotListAppend (x :: xs) xs' =
> rewrite odotListAppend xs xs' in
> rewrite odotAssociative x (odotList xs) (odotList xs') in Refl
------------------------------------
All four measure requirements hold:
------------------------------------
measPureSpec:
-------------
> -- |meas . pure `ExtEq` id|
> Framework.measPureSpec a = odotNeutrRight a
measJoinSpec:
-------------
< measJoinSpec : odotList . map odotList `ExtEq` odotList . concat
> Framework.measJoinSpec [] = Refl
> Framework.measJoinSpec ([] :: vss) =
> rewrite odotNeutrLeft (odotList (map odotList vss)) in
> rewrite measJoinSpec vss in
> Refl
> Framework.measJoinSpec ((v :: vs) :: vss) =
> rewrite sym (odotAssociative v (odotList vs) (odotList (map odotList vss))) in
> rewrite measJoinSpec vss in
> rewrite sym (odotListAppend vs (concat vss)) in
> Refl
>
> {--
> -- names
> let odot = odot in
> let odotL = odotList in
> -- proof contexts
> let c = odot v in
> let c' = \prf=> v `odot` (odotL vs `odot` prf) in
> -- proof steps
> let useOdotAssoc = odotAssociative v (odotL vs) (odotL (map odotL vss)) in
> let useIH = measJoinSpec vss in
> let useOdotAppend = odotListAppend vs (concat vss) in
> -------------------------------------------------------------------------
> ( ( odotL . map odotL) ((v :: vs) :: vss) ) ={ Refl }=
>
> ( (v `odot` odotL vs) `odot` (odotL (map odotL vss)) ) ={sym useOdotAssoc}=
>
> ( v `odot` (odotL vs `odot` (odotL (map odotL vss))) ) ={cong {f=c'} useIH}=
>
> ( v `odot` (odotL vs `odot` (odotL (concat vss))) ) ={cong {f=c} (sym useOdotAppend)}=
>
> ( v `odot` (odotL (vs ++ concat vss)) ) ={Refl}=
>
> ( (odotL . concat) ((v :: vs) :: vss) ) QED
> --}
measPlusSpec:
-------------
< measPlusSpec : (v : Val) -> (mv : List Val) -> NotEmpty mv ->
< odotList . map ((<+>) v) mv = (((<+>) v) . odotList) mv
> Framework.measPlusSpec _ [] ne = void ne
> Framework.measPlusSpec v (x :: []) _ = rewrite (odotNeutrRight x) in
> rewrite (odotNeutrRight (v <+> x)) in
> Refl
> Framework.measPlusSpec v (x :: (x' :: xs)) _ =
> rewrite measPlusSpec v (x' :: xs) () in
> rewrite sym (oplusOdotDistrLeft v x (odotList (x' :: xs))) in
> Refl
> {-
> -- proof steps
> let useIH = measPlusSpec v (x' :: xs) () in
> let useOplusOdotDistr = oplusOdotDistrLeft v x (odotList (x' :: xs)) in
> -------------------------------------------------------------------
>
> ( (meas . map (v <+>)) (x :: (x' :: xs)) ) ={ Refl }=
>
> ( odotList ((v <+> x) :: map (v <+>) (x' :: xs)) ) ={ Refl }=
>
> ( (v <+> x) `odot` odotList (map (v <+>) (x' :: xs)) ) ={ cong useIH }=
>
> ( (v <+> x) `odot` ((v <+>) . odotList) (x' :: xs) ) ={ sym useOplusOdotDistr }=
>
> ( v <+> (x `odot` odotList (x' :: xs) ) ) ={ Refl }=
>
> ( ((v <+>) . odotList) (x :: (x' :: xs)) ) QED
> --}
measMonSpec:
------------
If |odot| is monotone with respect to |<=|,
> odotMon : {a, b, c, d : Val} -> a <= b -> c <= d -> (a `odot` c) <= (b `odot` d)
we can moreover show
< measMonSpec : {A : Type} -> (f, g : A -> Val) -> ((a : A) -> (f a) <= (g a)) ->
< (ma : M A) -> meas (map f ma) <= meas (map g ma)
> Framework.measMonSpec f g ee [] = lteRefl