include "utilities/extralib.ma". include "common/IOMonad.ma". include "common/Integers.ma". include "common/Events.ma". include "common/Mem.ma". record execstep (outty:Type[0]) (inty:outty → Type[0]) : Type[1] ≝ { genv : Type[0] ; state : Type[0] ; is_final : state → option int ; mem_of_state : state → mem ; step : genv → state → IO outty inty (trace×state) }. let rec repeat (n:nat) (outty:Type[0]) (inty:outty → Type[0]) (exec:execstep outty inty) (g:genv ?? exec) (s:state ?? exec) : IO outty inty (trace × (state ?? exec)) ≝ match n with [ O ⇒ Value ??? 〈E0, s〉 | S n' ⇒ ! 〈t1,s1〉 ← step ?? exec g s; repeat n' ?? exec g s1 ]. let rec trace_map (A,B:Type[0]) (f:A → res (trace × B)) (l:list A) on l : res (trace × (list B)) ≝ match l with [ nil ⇒ OK ? 〈E0, [ ]〉 | cons h t ⇒ do 〈tr,h'〉 ← f h; do 〈tr',t'〉 ← trace_map … f t; OK ? 〈tr ⧺ tr',h'::t'〉 ]. (* A (possibly non-terminating) execution. *) coinductive execution (state:Type[0]) (output:Type[0]) (input:output → Type[0]) : Type[0] ≝ | e_stop : trace → int → mem → execution state output input | e_step : trace → state → execution state output input → execution state output input | e_wrong : errmsg → execution state output input | e_interact : ∀o:output. (input o → execution state output input) → execution state output input. (* This definition is slightly awkward because of the need to provide resumptions. It records the last trace/state passed in, then recursively processes the next state. *) let corec exec_inf_aux (output:Type[0]) (input:output → Type[0]) (exec:execstep output input) (ge:genv ?? exec) (s:IO output input (trace×(state ?? exec))) : execution ??? ≝ match s with [ Wrong m ⇒ e_wrong ??? m | Value v ⇒ match v with [ pair t s' ⇒ match is_final ?? exec s' with [ Some r ⇒ e_stop ??? t r (mem_of_state ?? exec s') | None ⇒ e_step ??? t s' (exec_inf_aux ??? ge (step ?? exec ge s')) ] ] | Interact out k' ⇒ e_interact ??? out (λv. exec_inf_aux ??? ge (k' v)) ]. lemma execution_cases: ∀o,i,s.∀e:execution s o i. e = match e with [ e_stop tr r m ⇒ e_stop ??? tr r m | e_step tr s e ⇒ e_step ??? tr s e | e_wrong m ⇒ e_wrong ??? m | e_interact o k ⇒ e_interact ??? o k ]. #o #i #s #e cases e; //; qed. axiom exec_inf_aux_unfold: ∀o,i,exec,ge,s. exec_inf_aux o i exec ge s = match s with [ Wrong m ⇒ e_wrong ??? m | Value v ⇒ match v with [ pair t s' ⇒ match is_final ?? exec s' with [ Some r ⇒ e_stop ??? t r (mem_of_state ?? exec s') | None ⇒ e_step ??? t s' (exec_inf_aux ??? ge (step ?? exec ge s')) ] ] | Interact out k' ⇒ e_interact ??? out (λv. exec_inf_aux ??? ge (k' v)) ]. (* #exec #ge #s >(execution_cases ? (exec_inf_aux …)) cases s [ #o #k | #x cases x #tr #s' (* XXX Can't unfold exec_inf_aux here *) | ] whd in ⊢ (??%%); //; qed. *) record fullexec (outty:Type[0]) (inty:outty → Type[0]) : Type[1] ≝ { es1 :> execstep outty inty ; program : Type[0] ; make_initial_state : program → res (genv ?? es1 × (state ?? es1)) }. definition exec_inf : ∀o,i.∀fx:fullexec o i. ∀p:program ?? fx. execution (state ?? fx) o i ≝ λo,i,fx,p. match make_initial_state ?? fx p with [ OK gs ⇒ exec_inf_aux ?? fx (\fst gs) (Value … 〈E0,\snd gs〉) | Error m ⇒ e_wrong ??? m ]. record execstep' (outty:Type[0]) (inty:outty → Type[0]) : Type[1] ≝ { es0 :> execstep outty inty ; initial : state ?? es0 → Prop }. alias symbol "and" (instance 2) = "logical and". record related_semantics : Type[1] ≝ { output : Type[0] ; input : output → Type[0] ; sem1 : execstep' output input ; sem2 : execstep' output input ; ge1 : genv ?? sem1 ; ge2 : genv ?? sem2 ; match_states : state ?? sem1 → state ?? sem2 → Prop ; match_initial_states : ∀st1. (initial ?? sem1) st1 → ∃st2. (initial ?? sem2) st2 ∧ match_states st1 st2 ; match_final_states : ∀st1,st2,r. match_states st1 st2 → (is_final ?? sem1) st1 = Some ? r → (is_final ?? sem2) st2 = Some ? r }. record simulation : Type[1] ≝ { sems :> related_semantics ; sim : ∀st1,st2,t,st1'. P_io' ?? (trace × (state ?? (sem1 sems))) (λr. r = 〈t,st1'〉) (step ?? (sem1 sems) (ge1 sems) st1) → match_states sems st1 st2 → ∃st2':(state ?? (sem2 sems)).(∃n:nat.P_io' ?? (trace × (state ?? (sem2 sems))) (λr. r = 〈t,st2'〉) (repeat n ?? (sem2 sems) (ge2 sems) st2)) ∧ match_states sems st1' st2' }.