include "joint/Joint.ma". include "utilities/bindLists.ma". (* inductive block_step (p : stmt_params) (globals : list ident) : Type[0] ≝ | block_seq : joint_seq p globals → block_step p globals | block_skip : label → block_step p globals. definition if_seq : ∀p,globals.∀A:Type[2].block_step p globals → A → A → A ≝ λp,g,A,s.match s with [ block_seq _ ⇒ λx,y.x | _ ⇒ λx,y.y ]. definition stmt_of_block_step : ∀p : stmt_params.∀globals. ∀s : block_step p globals.if_seq … s (succ p) unit → joint_statement p globals ≝ λp,g,s.match s return λx.if_seq ??? x ?? → joint_statement ?? with [ block_seq s' ⇒ λnxt.sequential … s' nxt | block_skip l ⇒ λ_.GOTO … l ]. definition seq_to_block_step_list : ∀p : stmt_params.∀globals. list (joint_seq p globals) → list (block_step p globals) ≝ λp,globals.map ?? (block_seq ??). coercion block_step_from_seq_list : ∀p : stmt_params.∀globals. ∀l:list (joint_seq p globals). list (block_step p globals) ≝ seq_to_block_step_list on _l:list (joint_seq ??) to list (block_step ??). definition is_inr : ∀A,B.A + B → bool ≝ λA,B,x.match x with [ inl _ ⇒ true | inr _ ⇒ false ]. definition is_inl : ∀A,B.A + B → bool ≝ λA,B,x.match x with [ inl _ ⇒ true | inr _ ⇒ false ]. definition skip_block ≝ λp,globals,A. (list (block_step p globals)) × A.*) (* move *) let rec bind_new_P R X (P : X → Prop) (b : bind_new R X) on b : Prop ≝ match b with [ bret x ⇒ P x | bnew f ⇒ ∀x.bind_new_P … P (f x) ]. definition BindNewP : ∀R.MonadPred (BindNew R) ≝ λR.mk_MonadPred (BindNew R) (bind_new_P R) ???. [ #X #P #x #K @K | #X #Y #Pin #Pout #m elim m -m [ #x #H #f #G @G @H | #g #IH #H #f #G #x @IH [ @H | @G ] ] | #X #P #Q #H #m elim m -m [ #x #Px @H @Px | #g #IH #Pg #x @IH @Pg ] ] qed. definition not_empty : ∀X.list X → Prop ≝ λX,l.match l with [ nil ⇒ False | _ ⇒ True ]. definition split_on_last : ∀X.list X → option ((list X) × X) ≝ λX.foldr … (λel,acc.match acc with [ None ⇒ Some ? 〈[ ], el〉 | Some pr ⇒ Some ? 〈el :: \fst pr, \snd pr〉 ]) (None ?). lemma split_on_last_inv : ∀X,l. match split_on_last X l with [ None ⇒ l = [ ] | Some pr ⇒ l = \fst pr @ [ \snd pr] ]. #X #l elim l -l normalize nodelta [ // ] #hd #tl change with (match split_on_last ? tl in option with [ _ ⇒ ?]) in match (split_on_last ? (hd :: tl)); cases (split_on_last ??) normalize nodelta [2: * #pref #last ] #EQ >EQ normalize % qed. lemma split_on_last_hit : ∀X,pref,last. split_on_last X (pref @ [last]) = Some ? 〈pref, last〉. #X #pref elim pref -pref normalize [//] #hd #tl #IH #last >IH % qed. (* the label input is to accomodate ERTLptr pass the postfix is for the case CALL ↦ instrs* ; CALL; instrs* *) definition step_block ≝ λp,globals.list (code_point p → joint_seq p globals) × (code_point p → joint_step p globals) × (list (joint_seq p globals)). definition fin_block ≝ λp,globals.(list (joint_seq p globals))×(joint_fin_step p). (*definition seq_to_skip_block : ∀p,g,A.seq_block p g A → skip_block p g A ≝ λp,g,A,b.〈\fst b, \snd b〉. coercion skip_from_seq_block : ∀p,g,A.∀b : seq_block p g A.skip_block p g A ≝ seq_to_skip_block on _b : seq_block ??? to skip_block ???.*) definition bind_step_block ≝ λp.λglobals. bind_new register (step_block p globals). unification hint 0 ≔ p, g; P ≟ step_block p g (*---------------------------------------*)⊢ bind_step_block p g ≡ bind_new register P. definition bind_fin_block ≝ λp : stmt_params.λglobals. bind_new register (fin_block p globals). unification hint 0 ≔ p : stmt_params, g; P ≟ fin_block p g (*---------------------------------------*)⊢ bind_fin_block p g ≡ bind_new register P. definition bind_seq_list ≝ λp : stmt_params.λglobals. bind_new register (list (joint_seq p globals)). unification hint 0 ≔ p : stmt_params, g; S ≟ joint_seq p g, L ≟ list S (*---------------------------------------*)⊢ bind_seq_list p g ≡ bind_new register L. definition add_dummy_variance : ∀X,Y : Type[0].list Y → list (X → Y) ≝ λX,Y.map … (λx.λ_.x). definition ensure_step_block : ∀p : params.∀globals. list (joint_seq p globals) → step_block p globals ≝ λp,g,l. match l return λ_.step_block ?? with [ nil ⇒ 〈[ ], λ_.NOOP ??, [ ]〉 | cons hd tl ⇒ 〈[ ], λ_.hd, tl〉 ]. coercion step_block_from_seq_list nocomposites : ∀p : params.∀g.∀l : list (joint_seq p g).step_block p g ≝ ensure_step_block on _l : list (joint_seq ??) to step_block ??. (*definition bind_skip_block ≝ λp : stmt_params.λglobals,A. bind_new (localsT p) (skip_block p globals A). unification hint 0 ≔ p : stmt_params, g, A; B ≟ skip_block p g A, R ≟ localsT p (*---------------------------------------*)⊢ bind_skip_block p g A ≡ bind_new R B. definition bind_seq_to_skip_block : ∀p,g,A.bind_seq_block p g A → bind_skip_block p g A ≝ λp,g,A.m_map ? (seq_block p g A) (skip_block p g A) (λx.x). coercion bind_skip_from_seq_block : ∀p,g,A.∀b:bind_seq_block p g A.bind_skip_block p g A ≝ bind_seq_to_skip_block on _b : bind_seq_block ??? to bind_skip_block ???.*) (* definition block_classifier ≝ λp,g.λb : other_block p g. seq_or_fin_step_classifier ?? (\snd b). *) (*definition seq_block_from_seq_list : ∀p : stmt_params.∀g.list (joint_seq p g) → seq_block p g (joint_step p g) ≝ λp,g,l.let 〈pre,post〉 ≝ split_on_last … (NOOP ??) l in 〈pre, (post : joint_step ??)〉. definition bind_seq_block_from_bind_seq_list : ∀p : stmt_params.∀g.bind_new (localsT p) (list (joint_seq p g)) → bind_seq_block p g (joint_step p g) ≝ λp.λg.m_map … (seq_block_from_seq_list …). definition bind_seq_block_step : ∀p,g.bind_seq_block p g (joint_step p g) → bind_seq_block p g (joint_step p g) + bind_seq_block p g (joint_fin_step p) ≝ λp,g.inl …. coercion bind_seq_block_from_step : ∀p,g.∀b:bind_seq_block p g (joint_step p g). bind_seq_block p g (joint_step p g) + bind_seq_block p g (joint_fin_step p) ≝ bind_seq_block_step on _b : bind_seq_block ?? (joint_step ??) to bind_seq_block ?? (joint_step ??) + bind_seq_block ?? (joint_fin_step ?). definition bind_seq_block_fin_step : ∀p,g.bind_seq_block p g (joint_fin_step p) → bind_seq_block p g (joint_step p g) + bind_seq_block p g (joint_fin_step p) ≝ λp,g.inr …. coercion bind_seq_block_from_fin_step : ∀p,g.∀b:bind_seq_block p g (joint_fin_step p). bind_seq_block p g (joint_step p g) + bind_seq_block p g (joint_fin_step p) ≝ bind_seq_block_fin_step on _b : bind_seq_block ?? (joint_fin_step ?) to bind_seq_block ?? (joint_step ??) + bind_seq_block ?? (joint_fin_step ?). definition seq_block_bind_seq_block : ∀p : stmt_params.∀g,A.seq_block p g A → bind_seq_block p g A ≝ λp,g,A.bret …. coercion seq_block_to_bind_seq_block : ∀p : stmt_params.∀g,A.∀b:seq_block p g A.bind_seq_block p g A ≝ seq_block_bind_seq_block on _b : seq_block ??? to bind_seq_block ???. definition joint_step_seq_block : ∀p : stmt_params.∀g.joint_step p g → seq_block p g (joint_step p g) ≝ λp,g,x.〈[ ], x〉. coercion joint_step_to_seq_block : ∀p : stmt_params.∀g.∀b : joint_step p g.seq_block p g (joint_step p g) ≝ joint_step_seq_block on _b : joint_step ?? to seq_block ?? (joint_step ??). definition joint_fin_step_seq_block : ∀p : stmt_params.∀g.joint_fin_step p → seq_block p g (joint_fin_step p) ≝ λp,g,x.〈[ ], x〉. coercion joint_fin_step_to_seq_block : ∀p : stmt_params.∀g.∀b : joint_fin_step p.seq_block p g (joint_fin_step p) ≝ joint_fin_step_seq_block on _b : joint_fin_step ? to seq_block ?? (joint_fin_step ?). definition seq_list_seq_block : ∀p:stmt_params.∀g.list (joint_seq p g) → seq_block p g (joint_step p g) ≝ λp,g,l.let pr ≝ split_on_last … (NOOP ??) l in 〈\fst pr, \snd pr〉. coercion seq_list_to_seq_block : ∀p:stmt_params.∀g.∀l:list (joint_seq p g).seq_block p g (joint_step p g) ≝ seq_list_seq_block on _l : list (joint_seq ??) to seq_block ?? (joint_step ??). definition bind_seq_list_bind_seq_block : ∀p:stmt_params.∀g.bind_new (localsT p) (list (joint_seq p g)) → bind_seq_block p g (joint_step p g) ≝ λp,g.m_map ??? (λx : list (joint_seq ??).(x : seq_block ???)). coercion bind_seq_list_to_bind_seq_block : ∀p:stmt_params.∀g.∀l:bind_new (localsT p) (list (joint_seq p g)).bind_seq_block p g (joint_step p g) ≝ bind_seq_list_bind_seq_block on _l : bind_new ? (list (joint_seq ??)) to bind_seq_block ?? (joint_step ??). *) notation > "x ~❨ B , l ❩~> y 'in' c" with precedence 56 for @{'block_in_code $c $x $B $l $y}. notation < "hvbox(x ~❨ B , break l ❩~> y \nbsp 'in' \nbsp break c)" with precedence 56 for @{'block_in_code $c $x $B $l $y}. notation > "x ~❨ B , l , r ❩~> y 'in' c" with precedence 56 for @{'bind_block_in_code $c $x $B $l $r $y}. notation < "hvbox(x ~❨ B , break l , break r ❩~> y \nbsp 'in' \nbsp break c)" with precedence 56 for @{'bind_block_in_code $c $x $B $l $r $y}. definition step_in_code ≝ λp,globals.λc : codeT p globals.λsrc : code_point p.λs : joint_step p globals. λdst : code_point p. ∃nxt.stmt_at … c src = Some ? (sequential … s nxt) ∧ point_of_succ … src nxt = dst. definition fin_step_in_code ≝ λp,globals.λc : codeT p globals.λsrc : code_point p.λs : joint_fin_step p. stmt_at … c src = Some ? (final … s). let rec seq_list_in_code p globals (c : codeT p globals) (src : code_point p) (B : list (joint_seq p globals)) (l : list (code_point p)) (dst : code_point p) on B : Prop ≝ match B with [ nil ⇒ l = [ ] ∧ src = dst | cons hd tl ⇒ ∃mid,rest.l = src :: rest ∧ step_in_code … c src hd mid ∧ seq_list_in_code … c mid tl rest dst ]. (* leaving it out because it can be misleading interpretation "step list in code" 'block_in_code c x B l y = (step_list_in_code ?? c x B l y). *) lemma seq_list_in_code_ne : ∀p,globals.∀c : codeT p globals.∀src,B,l,dst. not_empty ? B → seq_list_in_code … c src B l dst → ∃post.l = src :: post. #p #globals #c #src * [ #l #dst * ] #hd #tl #l #dst #_ * #mid * #post ** #EQl #_ #_ %{EQl} qed. (* #b lapply src -src elim b -b [ #src #l #dst * ] #hd * [2: #hd' #tl ] #IH #src * [1,3: #dst #_ * ] #mid * [ #dst #_ ** #EQ destruct #_ * |4: #mid' #tl #dst #_ ** #_ #_ * ] [2: #dst #_ ** #EQ destruct normalize nodelta #_ #_ %{[ ]} % ] #mid' #rest #dst #_ ** #EQ destruct #H1 #H2 cases (IH … I H2) #post normalize nodelta #EQ destruct %{(mid' :: post)} % qed. *) lemma seq_list_in_code_append : ∀p,globals.∀c : codeT p globals.∀src,B1,l1,mid,B2,l2,dst. seq_list_in_code … c src B1 l1 mid → seq_list_in_code … c mid B2 l2 dst → seq_list_in_code … c src (B1@B2) (l1@l2) dst. #p #globals #c #src #B1 lapply src -src elim B1 -B1 [ #src #l1 #mid #B2 #l2 #dst * #EQ1 #EQ2 destruct #H @H | #hd #tl #IH #src #l1 #mid #B2 #l2 #dst * #mid' * #rest ** #EQ destruct #H1 #H2 #H3 %{mid'} %{(rest@l2)} %{(IH … H2 H3)} %{H1} % ] qed. lemma seq_list_in_code_append_inv : ∀p,globals.∀c : codeT p globals.∀src,B1,B2,l,dst. seq_list_in_code … c src (B1@B2) l dst → ∃l1,mid,l2.l = l1 @ l2 ∧ seq_list_in_code … c src B1 l1 mid ∧ seq_list_in_code … c mid B2 l2 dst. #p #globals #c #src #B1 lapply src -src elim B1 -B1 [ #src #B2 #l #dst #H %{[ ]} %{src} %{l} %{H} % % % | #hd #tl #IH #src #B2 #l2 #dst * #mid * #rest ** #EQ destruct #H1 #H2 elim (IH … H2) #l1 * #mid' * #l2 ** #EQ destruct #G1 #G2 %{(src::l1)} %{mid'} %{l2} %{G2} %{(refl …)} % [| %[| %{G1} %{H1} % ]] ] qed. definition map_eval : ∀X,Y : Type[0].list (X → Y) → X → list Y ≝ λX,Y,l,x.map … (λf.f x) l. definition step_block_in_code ≝ λp,g.λc : codeT p g.λsrc.λb : step_block p g.λl,dst. let pref ≝ \fst (\fst b) in let stp ≝ \snd (\fst b) in let post ≝ \snd b in ∃l1,mid1,mid2,l2. l = l1 @ mid1 :: l2 ∧ seq_list_in_code ?? c src (map_eval … pref mid1) l1 mid1 ∧ step_in_code ?? c mid1 (stp mid1) mid2 ∧ seq_list_in_code ?? c mid2 post l2 dst. lemma map_compose : ∀X,Y,Z,f,g.∀l : list X.map Y Z f (map X Y g l) = map … (f∘g) l. #X #Y #Z #f #g #l elim l -l normalize // qed. lemma map_ext_eq : ∀X,Y,f,g.∀l : list X.(∀x.f x = g x) → map X Y f l = map X Y g l. #X #Y #f #g #l #H elim l -l normalize // qed. lemma map_id : ∀X.∀l : list X.map X X (λx.x) l = l. #X #l elim l -l normalize // qed. definition fin_block_in_code ≝ λp,g.λc:codeT p g.λsrc.λB : fin_block p g.λl.λdst : unit. ∃pref,mid.l = pref @ [mid] ∧ seq_list_in_code … c src (\fst B) pref mid ∧ fin_step_in_code … c mid (\snd B). interpretation "step block in code" 'block_in_code c x B l y = (step_block_in_code ?? c x B l y). interpretation "fin block in code" 'block_in_code c x B l y = (fin_block_in_code ?? c x B l y). lemma coerced_step_list_in_code : ∀p : params.∀g,c,src.∀b : list (joint_seq p g).∀l,dst. src ~❨b, l❩~> dst in c → match b with [ nil ⇒ step_in_code … c src (NOOP …) dst | _ ⇒ seq_list_in_code … c src b l dst ]. #p #g #c #src * [2: #hd #tl] #l #dst * #pref * #mid1 * #mid2 * #post *** #EQ * #EQ' #EQ'' #stp_in_code [ #tl_in | * #EQ''' #EQ''''] whd destruct [ % [| % [| %{tl_in} %{stp_in_code} % ]] | @stp_in_code ] qed. let rec bind_new_instantiates B X (x : X) (b : bind_new B X) (l : list B) on b : Prop ≝ match b with [ bret B ⇒ match l with [ nil ⇒ x = B | _ ⇒ False ] | bnew f ⇒ match l with [ nil ⇒ False | cons r l' ⇒ bind_new_instantiates B X x (f r) l' ] ]. lemma bind_new_instantiates_bind_inv : ∀B,X,Y,y. ∀b : bind_new B X.∀f : X → bind_new B Y.∀l. bind_new_instantiates B Y y (!y ← b; f y) l → ∃x,l1,l2.l = l1@l2 ∧ bind_new_instantiates B X x b l1 ∧ bind_new_instantiates B Y y (f x) l2. #B #X #Y #y #b elim b -b [ #x #f #l >m_return_bind #H %{x} %{[ ]} %{l} %{H} %% | #g #IH #f * [ * ] #hd #tl whd in ⊢ (%→?); #H cases (IH … H) #x * #l1 * #l2 ** #EQ #H1 #H2 %{x} %{(hd::l1)} %{l2} %{H2} %{H1} >EQ % ] qed. lemma bind_new_instantiates_bind : ∀B,X,Y,x,y. ∀b : bind_new B X.∀f : X → bind_new B Y.∀l1,l2. bind_new_instantiates B X x b l1 → bind_new_instantiates B Y y (f x) l2 → bind_new_instantiates B Y y (!y ← b; f y) (l1@l2). #B #X #Y #x #y #b elim b -b [ #x' #f * [2: #hd #tl ] #l2 * #H @H | #g #IH #f * [ #l2 * ] #hd #tl #l2 #H1 #H2 whd @IH assumption ] qed. definition bind_step_block_in_code ≝ λp,g.λc:codeT p g.λsrc.λB : bind_step_block p g.λlbls,regs.λdst. ∃b.bind_new_instantiates … b B regs ∧ src ~❨b, lbls❩~> dst in c. definition bind_fin_block_in_code ≝ λp,g.λc:codeT p g.λsrc.λB : bind_fin_block p g.λlbls,regs.λdst. ∃b.bind_new_instantiates … b B regs ∧ src ~❨b, lbls❩~> dst in c. interpretation "bound step block in code" 'bind_block_in_code c x B l r y = (bind_step_block_in_code ?? c x B l r y). interpretation "bound fin block in code" 'bind_block_in_code c x B l r y = (bind_fin_block_in_code ?? c x B l r y). (* generates ambiguity even if it shouldn't interpretation "seq block step in code" 'block_in_code c x B l y = (seq_block_step_in_code ?? c x B l y). interpretation "seq block fin step in code" 'block_in_code c x B l y = (seq_block_fin_step_in_code ?? c x B l y). *) (* definition seq_block_append : ∀p,g. ∀b1 : Σb.is_safe_block p g b. ∀b2 : seq_block p g. seq_block p g ≝ λp,g,b1,b2. 〈match b1 with [ mk_Sig instr prf ⇒ match \snd instr return λx.bool_to_Prop (is_inl … x) ∧ seq_or_fin_step_classifier … x = ? → ? with [ inl i ⇒ λprf.\fst b1 @ i :: \fst b2 | inr _ ⇒ λprf.⊥ ] prf ],\snd b2〉. cases prf #H1 #H2 assumption qed. definition other_block_append : ∀p,g. (Σb.block_classifier ?? b = cl_other) → other_block p g → other_block p g ≝ λp,g,b1,b2. 〈\fst b1 @ «\snd b1, pi2 … b1» :: \fst b2, \snd b2〉. definition seq_block_cons : ∀p : stmt_params.∀g.(Σs.step_classifier p g s = cl_other) → seq_block p g → seq_block p g ≝ λp,g,x,b.〈x :: \fst b,\snd b〉. definition other_block_cons : ∀p,g. (Σs.seq_or_fin_step_classifier p g s = cl_other) → other_block p g → other_block p g ≝ λp,g,x,b.〈x :: \fst b,\snd b〉. interpretation "seq block cons" 'cons x b = (seq_block_cons ?? x b). interpretation "other block cons" 'vcons x b = (other_block_cons ?? x b). interpretation "seq block append" 'append b1 b2 = (seq_block_append ?? b1 b2). interpretation "other block append" 'vappend b1 b2 = (other_block_append ?? b1 b2). definition step_to_block : ∀p,g. seq_or_fin_step p g → seq_block p g ≝ λp,g,s.〈[ ], s〉. coercion block_from_step : ∀p,g.∀s : seq_or_fin_step p g. seq_block p g ≝ step_to_block on _s : seq_or_fin_step ?? to seq_block ??. definition bind_seq_block_cons : ∀p : stmt_params.∀g,is_seq. (Σs.step_classifier p g s = cl_other) → bind_seq_block p g is_seq → bind_seq_block p g is_seq ≝ λp,g,is_seq,x.m_map ??? (λb.〈x::\fst b,\snd b〉). definition bind_other_block_cons : ∀p,g.(Σs.seq_or_fin_step_classifier p g s = cl_other) → bind_other_block p g → bind_other_block p g ≝ λp,g,x.m_map … (other_block_cons … x). let rec bind_pred_aux B X (P : X → Prop) (c : bind_new B X) on c : Prop ≝ match c with [ bret x ⇒ P x | bnew f ⇒ ∀b.bind_pred_aux B X P (f b) ]. let rec bind_pred_inj_aux B X (P : X → Prop) (c : bind_new B X) on c : bind_pred_aux B X P c → bind_new B (Σx.P x) ≝ match c return λx.bind_pred_aux B X P x → bind_new B (Σx.P x) with [ bret x ⇒ λprf.return «x, prf» | bnew f ⇒ λprf.bnew … (λx.bind_pred_inj_aux B X P (f x) (prf x)) ]. definition bind_pred ≝ λB. mk_InjMonadPred (BindNew B) (mk_MonadPred ? (bind_pred_aux B) ???) (λX,P,c_sig.bind_pred_inj_aux B X P c_sig (pi2 … c_sig)) ?. [ #X #P #Q #H #y elim y [ #x @H | #f #IH #G #b @IH @G] | #X #Y #Pin #Pout #m elim m [#x | #f #IH] #H #g #G [ @G @H | #b @(IH … G) @H] | #X #P #x #Px @Px | #X #P * #m elim m [#x | #f #IH] #H [ % | @bnew_proper #b @IH] ] qed. definition bind_seq_block_append : ∀p,g,is_seq.(Σb : bind_seq_block p g true.bind_pred ? (λb.step_classifier p g (\snd b) = cl_other) b) → bind_seq_block p g is_seq → bind_seq_block p g is_seq ≝ λp,g,is_seq,b1,b2. !«p, prf» ← mp_inject … b1; !〈post, last〉 ← b2; return 〈\fst p @ «\snd p, prf» :: post, last〉. definition bind_other_block_append : ∀p,g.(Σb : bind_other_block p g.bind_pred ? (λx.block_classifier ?? x = cl_other) b) → bind_other_block p g → bind_other_block p g ≝ λp,g,b1.m_bin_op … (other_block_append ??) (mp_inject … b1). interpretation "bind seq block cons" 'cons x b = (bind_seq_block_cons ??? x b). interpretation "bind other block cons" 'vcons x b = (bind_other_block_cons ?? x b). interpretation "bind seq block append" 'append b1 b2 = (bind_seq_block_append ??? b1 b2). interpretation "bind other block append" 'vappend b1 b2 = (bind_other_block_append ?? b1 b2). let rec instantiates_to B X (b : bind_new B X) (l : list B) (x : X) on b : Prop ≝ match b with [ bret B ⇒ match l with [ nil ⇒ x = B | _ ⇒ False ] | bnew f ⇒ match l with [ nil ⇒ False | cons r l' ⇒ instantiates_to B X (f r) l' x ] ]. lemma instantiates_to_bind_pred : ∀B,X,P,b,l,x.instantiates_to B X b l x → bind_pred B P b → P x. #B #X #P #b elim b [ #x * [ #y #EQ >EQ normalize // | #hd #tl #y *] | #f #IH * [ #y * | #hd #tl normalize #b #H #G @(IH … H) @G ] ] qed. lemma seq_block_append_proof_irr : ∀p,g,b1,b1',b2.pi1 ?? b1 = pi1 ?? b1' → seq_block_append p g b1 b2 = seq_block_append p g b1' b2. #p #g * #b1 #b1prf * #b1' #b1prf' #b2 #EQ destruct(EQ) % qed. lemma other_block_append_proof_irr : ∀p,g,b1,b1',b2.pi1 ?? b1 = pi1 ?? b1' → other_block_append p g b1 b2 = other_block_append p g b1' b2. #p #g * #b1 #b1prf * #b1' #b1prf' #b2 #EQ destruct(EQ) % qed. (* lemma is_seq_block_instance_append : ∀p,g,is_seq. ∀B1. ∀B2 : bind_seq_block p g is_seq. ∀l1,l2. ∀b1 : Σb.is_safe_block p g b. ∀b2 : seq_block p g. instantiates_to ? (seq_block p g) B1 l1 (pi1 … b1) → instantiates_to ? (seq_block p g) B2 l2 b2 → instantiates_to ? (seq_block p g) (B1 @ B2) (l1 @ l2) (b1 @ b2). #p #g * #B1 elim B1 -B1 [ #B1 | #f1 #IH1 ] #B1prf whd in B1prf; #B2 * [2,4: #r1 #l1' ] #l2 #b1 #b2 [1,4: *] whd in ⊢ (%→?); [ @IH1 | #EQ destruct(EQ) lapply b2 -b2 lapply l2 -l2 elim B2 -B2 [ #B2 | #f2 #IH2] * [2,4: #r2 #l2'] #b2 [1,4: *] whd in ⊢ (%→?); [ @IH2 | #EQ' whd destruct @seq_block_append_proof_irr % ] ] qed. lemma is_other_block_instance_append : ∀p,g. ∀B1 : Σb.bind_pred ? (λx.block_classifier p g x = cl_other) b. ∀B2 : bind_other_block p g. ∀l1,l2. ∀b1 : Σb.block_classifier p g b = cl_other. ∀b2 : other_block p g. instantiates_to ? (other_block p g) B1 l1 (pi1 … b1) → instantiates_to ? (other_block p g) B2 l2 b2 → instantiates_to ? (other_block p g) (B1 @@ B2) (l1 @ l2) (b1 @@ b2). #p #g * #B1 elim B1 -B1 [ #B1 | #f1 #IH1 ] #B1prf whd in B1prf; #B2 * [2,4: #r1 #l1' ] #l2 #b1 #b2 [1,4: *] whd in ⊢ (%→?); [ @IH1 | #EQ destruct(EQ) lapply b2 -b2 lapply l2 -l2 elim B2 -B2 [ #B2 | #f2 #IH2] * [2,4: #r2 #l2'] #b2 [1,4: *] whd in ⊢ (%→?); [ @IH2 | #EQ' whd destruct @other_block_append_proof_irr % ] ] qed. lemma other_fin_step_has_one_label : ∀p,g.∀s:(Σs.fin_step_classifier p g s = cl_other). match fin_step_labels ?? s with [ nil ⇒ False | cons _ tl ⇒ match tl with [ nil ⇒ True | _ ⇒ False ] ]. #p #g ** [#lbl || #ext] normalize [3: cases (ext_fin_step_flows p ext) [* [2: #lbl' * [2: #lbl'' #tl']]] normalize nodelta ] #EQ destruct % qed. definition label_of_other_fin_step : ∀p,g. (Σs.fin_step_classifier p g s = cl_other) → label ≝ λp,g,s. match fin_step_labels p ? s return λx.match x with [ nil ⇒ ? | cons _ tl ⇒ ?] → ? with [ nil ⇒ Ⓧ | cons lbl tl ⇒ λ_.lbl ] (other_fin_step_has_one_label p g s). (* definition point_seq_transition : ∀p,g.codeT p g → code_point p → code_point p → Prop ≝ λp,g,c,src,dst. match stmt_at … c src with [ Some stmt ⇒ match stmt with [ sequential sq nxt ⇒ point_of_succ … src nxt = dst | final fn ⇒ match fin_step_labels … fn with [ nil ⇒ False | cons lbl tl ⇒ match tl with [ nil ⇒ point_of_label … c lbl = Some ? dst | _ ⇒ False ] ] ] | None ⇒ False ]. lemma point_seq_transition_label_of_other_fin_step : ∀p,c,src.∀s : (Σs.fin_step_classifier p s = cl_other).∀dst. stmt_at ?? c src = Some ? s → point_seq_transition p c src dst → point_of_label … c (label_of_other_fin_step p s) = Some ? dst. #p #c #src ** [#lbl || #ext] #EQ1 #dst #EQ2 whd in match point_seq_transition; normalize nodelta >EQ2 normalize nodelta whd in ⊢ (?→??(????%)?); [#H @H | * ] lapply (other_fin_step_has_one_label ? «ext,?») cases (fin_step_labels p ? ext) normalize nodelta [*] #hd * normalize nodelta [2: #_ #_ *] * #H @H qed. lemma point_seq_transition_succ : ∀p,c,src.∀s,nxt.∀dst. stmt_at ?? c src = Some ? (sequential ?? s nxt) → point_seq_transition p c src dst → point_of_succ … src nxt = dst. #p #c #src #s #nxt #dst #EQ whd in match point_seq_transition; normalize nodelta >EQ normalize nodelta #H @H qed. *) definition if_other : ∀p,g.∀A : Type[2].seq_or_fin_step p g → A → A → A ≝ λp,g,A,c.match seq_or_fin_step_classifier p g c with [ cl_other ⇒ λx,y.x | _ ⇒ λx,y.y ]. definition other_step_in_code ≝ λp,g. λc : codeT p g. λsrc : code_point p. λs : seq_or_fin_step p g. match s return λx.if_other p g ? x (code_point p) unit → Prop with [ inl s'' ⇒ λdst.∃n.stmt_at … c src = Some ? (sequential … s'' n) ∧ ? | inr s'' ⇒ λdst.stmt_at … c src = Some ? (final … s'') ∧ ? ]. [ whd in dst; cases (seq_or_fin_step_classifier ???) in dst; normalize nodelta [1,2,3: #_ @True |*: #dst @(point_of_succ … src n = dst)] | whd in dst; lapply dst -dst lapply (refl … (seq_or_fin_step_classifier ?? (inr … s''))) cases (seq_or_fin_step_classifier ?? (inr … s'')) in ⊢ (???%→%); normalize nodelta [1,2,3: #_ #_ @True |*: #EQ #dst @(point_of_label … c (label_of_other_fin_step p g «s'', EQ») = Some ? dst) ] ] qed. definition if_other_sig : ∀p,g.∀B,C : Type[0].∀s : Σs.seq_or_fin_step_classifier p g s = cl_other. if_other p g ? s B C → B ≝ λp,g,B,C.?. ** #s whd in match (if_other ??????); cases (seq_or_fin_step_classifier ???) normalize nodelta #EQ destruct(EQ) #x @x qed. definition if_other_block_sig : ∀p,g.∀B,C : Type[0].∀b : Σb.block_classifier p g b = cl_other. if_other p g ? (\snd b) B C → B ≝ λp,g,B,C.?. ** #l #s #prf #x @(if_other_sig ???? «s, prf» x) qed. coercion other_sig_to_if nocomposites: ∀p,g.∀B,C : Type[0].∀s : Σs.seq_or_fin_step_classifier p g s = cl_other. ∀x : if_other p g ? s B C.B ≝ if_other_sig on _x : if_other ?? Type[0] ??? to ?. coercion other_block_sig_to_if nocomposites: ∀p,g.∀B,C : Type[0].∀s : Σs.block_classifier p g s = cl_other. ∀x : if_other p g ? (\snd s) B C.B ≝ if_other_block_sig on _x : if_other ?? Type[0] (\snd ?) ?? to ?. let rec other_list_in_code p g (c : codeT p g) src (b : list (Σs.seq_or_fin_step_classifier p g s = cl_other)) dst on b : Prop ≝ match b with [ nil ⇒ src = dst | cons hd tl ⇒ ∃mid. other_step_in_code p g c src hd mid ∧ other_list_in_code p g c mid tl dst ]. notation > "x ~❨ B ❩~> y 'in' c" with precedence 56 for @{'block_in_code $c $x $B $y}. notation < "hvbox(x ~❨ B ❩~> y \nbsp 'in' \nbsp break c)" with precedence 56 for @{'block_in_code $c $x $B $y}. interpretation "list in code" 'block_in_code c x B y = (other_list_in_code ?? c x B y). definition other_block_in_code : ∀p,g.codeT p g → code_point p → ∀b : other_block p g. if_other … (\snd b) (code_point p) unit → Prop ≝ λp,g,c,src,b,dst. ∃mid.src ~❨\fst b❩~> mid in c ∧ other_step_in_code p g c mid (\snd b) dst. interpretation "block in code" 'block_in_code c x B y = (other_block_in_code ?? c x B y). lemma other_list_in_code_append : ∀p,g.∀c : codeT p g. ∀x.∀b1 : list ?. ∀y.∀b2 : list ?.∀z. x ~❨b1❩~> y in c→ y ~❨b2❩~> z in c → x ~❨b1@b2❩~> z in c. #p#g#c#x#b1 lapply x -x elim b1 [2: ** #hd #hd_prf #tl #IH] #x #y #b2 #z [3: #EQ normalize in EQ; destruct #H @H] * #mid * normalize nodelta [ *#n ] #H1 #H2 #H3 whd normalize nodelta %{mid} %{(IH … H2 H3)} [ %{n} ] @H1 qed. lemma other_block_in_code_append : ∀p,g.∀c : codeT p g.∀x. ∀B1 : Σb.block_classifier p g b = cl_other. ∀y. ∀B2 : other_block p g. ∀z. x ~❨B1❩~> y in c → y ~❨B2❩~> z in c → x ~❨B1@@B2❩~> z in c. #p#g#c #x ** #hd1 *#tl1 #tl1prf #y * #hd2 #tl2 #z * #mid1 * #H1 #H2 * #mid2 * #G1 #G2 %{mid2} %{G2} whd in match (\fst ?); @(other_list_in_code_append … H1) %{y} %{H2 G1} qed. (* definition instr_block_in_function : ∀p : evaluation_params.∀fn : joint_internal_function (globals p) p. code_point p → ∀b : bind_other_block p. ? → Prop ≝ λp,fn,src,B,dst. ∃vars,B'. All ? (In ? (joint_if_locals … fn)) vars ∧ instantiates_to … B vars B' ∧ src ~❨B'❩~> dst in joint_if_code … fn. interpretation "bind block in function" 'block_in_code fn x B y = (instr_block_in_function ? fn x B y). lemma instr_block_in_function_trans : ∀p,fn,src. ∀B1 : ΣB.bind_pred ? (λb.block_classifier p b = cl_other) B. ∀mid. ∀B2 : bind_other_block p. ∀dst. src ~❨B1❩~> Some ? mid in fn → mid ~❨B2❩~> dst in fn → src ~❨B1@@B2❩~> dst in fn. #p#fn#src*#B1#B1prf#mid#B2#dst * #vars1 * #b1 ** #vars1_ok #b1B1 #b1_in * #vars2 * #b2 ** #vars2_ok #b2B2 #b2_in %{(vars1@vars2)} %{(«b1,instantiates_to_bind_pred … b1B1 B1prf» @@ b2)} /4 by All_append, conj, is_other_block_instance_append, other_block_in_code_append/ qed. *) *) *)