]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/turing.ma
a8290de51d7d32eb365fa3b7fdc9786b3bc2a8bb
[helm.git] / matita / matita / lib / turing / turing.ma
1 include "turing/mono.ma".
2 include "basics/vectors.ma".
3
4 (* We do not distinuish an input tape *)
5
6 (* tapes_no = number of ADDITIONAL working tapes *)
7
8 record mTM (sig:FinSet) (tapes_no:nat) : Type[1] ≝ 
9 { states : FinSet;
10   trans : states × (Vector (option sig) (S tapes_no)) → 
11     states  × (Vector (option (sig × move))(S tapes_no));
12   start: states;
13   halt : states → bool
14 }.
15
16 record mconfig (sig,states:FinSet) (n:nat): Type[0] ≝
17 { cstate : states;
18   ctapes : Vector (tape sig) (S n)
19 }.
20
21 lemma mconfig_expand: ∀sig,n,Q,c. 
22   c = mk_mconfig sig Q n (cstate ??? c) (ctapes ??? c).
23 #sig #n #Q * // 
24 qed.
25   
26 lemma mconfig_eq : ∀sig,n,M,c1,c2.
27   cstate sig n M c1 = cstate sig n M c2 → 
28     ctapes sig n M c1 = ctapes sig n M c2 →  c1 = c2.
29 #sig #n #M1 * #s1 #t1 * #s2 #t2 //
30 qed.
31
32 definition current_chars ≝ λsig.λn.λtapes.
33   vec_map ?? (current sig) (S n) tapes.
34
35 definition step ≝ λsig.λn.λM:mTM sig n.λc:mconfig sig (states ?? M) n.
36   let 〈news,mvs〉 ≝ trans sig n M 〈cstate ??? c,current_chars ?? (ctapes ??? c)〉 in
37   mk_mconfig ??? 
38     news 
39     (pmap_vec ??? (tape_move sig) ? (ctapes ??? c) mvs).
40
41 definition empty_tapes ≝ λsig.λn.
42 mk_Vector ? n (make_list (tape sig) (niltape sig) n) ?.
43 elim n // normalize //
44 qed.
45
46 (************************** Realizability *************************************)
47 definition loopM ≝ λsig,n.λM:mTM sig n.λi,cin.
48   loop ? i (step sig n M) (λc.halt sig n M (cstate ??? c)) cin.
49
50 lemma loopM_unfold : ∀sig,n,M,i,cin.
51   loopM sig n M i cin = loop ? i (step sig n M) (λc.halt sig n M (cstate ??? c)) cin.
52 // qed.
53
54 definition initc ≝ λsig,n.λM:mTM sig n.λtapes.
55   mk_mconfig sig (states sig n M) n (start sig n M) tapes.
56
57 definition Realize ≝ λsig,n.λM:mTM sig n.λR:relation (Vector (tape sig) ?).
58 ∀t.∃i.∃outc.
59   loopM sig n M i (initc sig n M t) = Some ? outc ∧ R t (ctapes ??? outc).
60
61 definition WRealize ≝ λsig,n.λM:mTM sig n.λR:relation (Vector (tape sig) ?).
62 ∀t,i,outc.
63   loopM sig n M i (initc sig n M t) = Some ? outc → R t (ctapes ??? outc).
64
65 definition Terminate ≝ λsig,n.λM:mTM sig n.λt. ∃i,outc.
66   loopM sig n M i (initc sig n M t) = Some ? outc.
67   
68 (* notation "M \vDash R" non associative with precedence 45 for @{ 'models $M $R}. *)
69 interpretation "multi realizability" 'models M R = (Realize ?? M R).
70
71 (* notation "M \VDash R" non associative with precedence 45 for @{ 'wmodels $M $R}. *)
72 interpretation "weak multi realizability" 'wmodels M R = (WRealize ?? M R).
73
74 interpretation "multi termination" 'fintersects M t = (Terminate ?? M t).
75
76 lemma WRealize_to_Realize : ∀sig,n .∀M: mTM sig n.∀R.
77   (∀t.M ↓ t) → M ⊫ R → M ⊨ R.
78 #sig #n #M #R #HT #HW #t cases (HT … t) #i * #outc #Hloop 
79 @(ex_intro … i) @(ex_intro … outc) % // @(HW … i) //
80 qed.
81
82 theorem Realize_to_WRealize : ∀sig,n.∀M:mTM sig n.∀R.
83   M ⊨ R → M ⊫ R.
84 #sig #n #M #R #H1 #inc #i #outc #Hloop 
85 cases (H1 inc) #k * #outc1 * #Hloop1 #HR >(loop_eq … Hloop Hloop1) //
86 qed.
87
88 definition accRealize ≝ λsig,n.λM:mTM sig n.λacc:states sig n M.λRtrue,Rfalse.
89 ∀t.∃i.∃outc.
90   loopM sig n M i (initc sig n M t) = Some ? outc ∧
91     (cstate ??? outc = acc → Rtrue t (ctapes ??? outc)) ∧ 
92     (cstate ??? outc ≠ acc → Rfalse t (ctapes ??? outc)).
93     
94 (* notation "M ⊨ [q: R1,R2]" non associative with precedence 45 for @{ 'cmodels $M $q $R1 $R2}. *)
95 interpretation "conditional multi realizability" 'cmodels M q R1 R2 = (accRealize ?? M q R1 R2).
96
97 (*************************** guarded realizablity *****************************)
98 definition GRealize ≝ λsig,n.λM:mTM sig n.
99  λPre:Vector (tape sig) ? →Prop.λR:relation (Vector (tape sig) ?).
100   ∀t.Pre t → ∃i.∃outc.
101    loopM sig n M i (initc sig n M t) = Some ? outc ∧ R t (ctapes ??? outc).
102   
103 definition accGRealize ≝ λsig,n.λM:mTM sig n.λacc:states sig n M.
104 λPre: Vector (tape sig) ? → Prop.λRtrue,Rfalse.
105 ∀t.Pre t → ∃i.∃outc.
106   loopM sig n M i (initc sig n M t) = Some ? outc ∧
107     (cstate ??? outc = acc → Rtrue t (ctapes ??? outc)) ∧ 
108     (cstate ??? outc ≠ acc → Rfalse t (ctapes ??? outc)).
109     
110 lemma WRealize_to_GRealize : ∀sig,n.∀M: mTM sig n.∀Pre,R.
111   (∀t.Pre t → M ↓ t) → M ⊫ R → GRealize sig n M Pre R.
112 #sig #n #M #Pre #R #HT #HW #t #HPre cases (HT … t HPre) #i * #outc #Hloop 
113 @(ex_intro … i) @(ex_intro … outc) % // @(HW … i) //
114 qed.
115
116 lemma Realize_to_GRealize : ∀sig,n.∀M: mTM sig n.∀P,R. 
117   M ⊨ R → GRealize sig n M P R.
118 #alpha #n #M #Pre #R #HR #t #HPre
119 cases (HR t) -HR #k * #outc * #Hloop #HR 
120 @(ex_intro ?? k) @(ex_intro ?? outc) % 
121   [ @Hloop | @HR ]
122 qed.
123
124 lemma acc_Realize_to_acc_GRealize: ∀sig,n.∀M:mTM sig n.∀q:states sig n M.∀P,R1,R2. 
125   M ⊨ [q:R1,R2] → accGRealize sig n M q P R1 R2.
126 #alpha #n #M #q #Pre #R1 #R2 #HR #t #HPre
127 cases (HR t) -HR #k * #outc * * #Hloop #HRtrue #HRfalse 
128 @(ex_intro ?? k) @(ex_intro ?? outc) % 
129   [ % [@Hloop] @HRtrue | @HRfalse]
130 qed.
131
132 (******************************** monotonicity ********************************)
133 lemma Realize_to_Realize : ∀sig,n.∀M:mTM sig n.∀R1,R2.
134   R1 ⊆ R2 → M ⊨ R1 → M ⊨ R2.
135 #alpha #n #M #R1 #R2 #Himpl #HR1 #intape
136 cases (HR1 intape) -HR1 #k * #outc * #Hloop #HR1
137 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
138 qed.
139
140 lemma WRealize_to_WRealize: ∀sig,n.∀M:mTM sig n.∀R1,R2.
141   R1 ⊆ R2 → WRealize sig n M R1 → WRealize sig n M R2.
142 #alpha #n #M #R1 #R2 #Hsub #HR1 #intape #i #outc #Hloop
143 @Hsub @(HR1 … i) @Hloop
144 qed.
145
146 lemma GRealize_to_GRealize : ∀sig,n.∀M:mTM sig n.∀P,R1,R2.
147   R1 ⊆ R2 → GRealize sig n M P R1 → GRealize sig n M P R2.
148 #alpha #n #M #P #R1 #R2 #Himpl #HR1 #intape #HP
149 cases (HR1 intape HP) -HR1 #k * #outc * #Hloop #HR1
150 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
151 qed.
152
153 lemma GRealize_to_GRealize_2 : ∀sig,n.∀M:mTM sig n.∀P1,P2,R1,R2.
154   P2 ⊆ P1 → R1 ⊆ R2 → GRealize sig n M P1 R1 → GRealize sig n M P2 R2.
155 #alpha #n #M #P1 #P2 #R1 #R2 #Himpl1 #Himpl2 #H1 #intape #HP
156 cases (H1 intape (Himpl1 … HP)) -H1 #k * #outc * #Hloop #H1
157 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
158 qed.
159
160 lemma acc_Realize_to_acc_Realize: ∀sig,n.∀M:mTM sig n.∀q:states sig n M.
161  ∀R1,R2,R3,R4. 
162   R1 ⊆ R3 → R2 ⊆ R4 → M ⊨ [q:R1,R2] → M ⊨ [q:R3,R4].
163 #alpha #n #M #q #R1 #R2 #R3 #R4 #Hsub13 #Hsub24 #HRa #intape
164 cases (HRa intape) -HRa #k * #outc * * #Hloop #HRtrue #HRfalse 
165 @(ex_intro ?? k) @(ex_intro ?? outc) % 
166   [ % [@Hloop] #Hq @Hsub13 @HRtrue // | #Hq @Hsub24 @HRfalse //]
167 qed.
168
169 (**************************** A canonical relation ****************************)
170
171 definition R_mTM ≝ λsig,n.λM:mTM sig n.λq.λt1,t2.
172 ∃i,outc.
173   loopM ? n M i (mk_mconfig ??? q t1) = Some ? outc ∧ 
174   t2 = (ctapes ??? outc).
175   
176 lemma R_mTM_to_R: ∀sig,n.∀M:mTM sig n.∀R. ∀t1,t2. 
177   M ⊫ R → R_mTM ?? M (start sig n M) t1 t2 → R t1 t2.
178 #sig #n #M #R #t1 #t2 whd in ⊢ (%→?); #HMR * #i * #outc *
179 #Hloop #Ht2 >Ht2 @(HMR … Hloop)
180 qed.
181
182 (******************************** NOP Machine *********************************)
183
184 (* NO OPERATION
185    t1 = t2 
186   
187 definition nop_states ≝ initN 1.
188 definition start_nop : initN 1 ≝ mk_Sig ?? 0 (le_n … 1). *)
189
190 definition nop ≝ 
191   λalpha:FinSet.λn.mk_mTM alpha n nop_states
192   (λp.let 〈q,a〉 ≝ p in 〈q,mk_Vector ? (S n) (make_list ? (None ?) (S n)) ?〉)
193   start_nop (λ_.true).
194 elim n normalize //
195 qed.
196   
197 definition R_nop ≝ λalpha,n.λt1,t2:Vector (tape alpha) (S n).t2 = t1.
198
199 lemma sem_nop :
200   ∀alpha,n.nop alpha n⊨ R_nop alpha n.
201 #alpha #n #intapes @(ex_intro ?? 1) 
202 @(ex_intro … (mk_mconfig ??? start_nop intapes)) % % 
203 qed.
204
205 lemma nop_single_state: ∀sig,n.∀q1,q2:states ? n (nop sig n). q1 = q2.
206 normalize #sig #n0 * #n #ltn1 * #m #ltm1 
207 generalize in match ltn1; generalize in match ltm1;
208 <(le_n_O_to_eq … (le_S_S_to_le … ltn1)) <(le_n_O_to_eq … (le_S_S_to_le … ltm1)) 
209 // qed.
210
211 (************************** Sequential Composition ****************************)
212 definition null_action ≝ λsig.λn.
213 mk_Vector ? (S n) (make_list (option (sig × move)) (None ?) (S n)) ?.
214 elim (S n) // normalize //
215 qed.
216
217 lemma tape_move_null_action: ∀sig,n,tapes.
218   pmap_vec ??? (tape_move sig) (S n) tapes (null_action sig n) = tapes.
219 #sig #n #tapes cases tapes -tapes #tapes whd in match (null_action ??);
220 #Heq @Vector_eq <Heq -Heq elim tapes //
221 #a #tl #Hind whd in ⊢ (??%?); @eq_f2 // @Hind
222 qed.
223
224 definition seq_trans ≝ λsig,n. λM1,M2 : mTM sig n. 
225 λp. let 〈s,a〉 ≝ p in
226   match s with 
227   [ inl s1 ⇒ 
228       if halt sig n M1 s1 then 〈inr … (start sig n M2), null_action sig n〉
229       else let 〈news1,m〉 ≝ trans sig n M1 〈s1,a〉 in 〈inl … news1,m〉
230   | inr s2 ⇒ let 〈news2,m〉 ≝ trans sig n M2 〈s2,a〉 in 〈inr … news2,m〉
231   ].
232  
233 definition seq ≝ λsig,n. λM1,M2 : mTM sig n. 
234   mk_mTM sig n
235     (FinSum (states sig n M1) (states sig n M2))
236     (seq_trans sig n M1 M2) 
237     (inl … (start sig n M1))
238     (λs.match s with 
239       [ inl _ ⇒ false | inr s2 ⇒ halt sig n M2 s2]). 
240
241 (* notation "a · b" right associative with precedence 65 for @{ 'middot $a $b}. *)
242 interpretation "sequential composition" 'middot a b = (seq ?? a b).
243
244 definition lift_confL ≝ 
245   λsig,n,S1,S2,c.match c with 
246   [ mk_mconfig s t ⇒ mk_mconfig sig (FinSum S1 S2) n (inl … s) t ].
247   
248 definition lift_confR ≝ 
249   λsig,n,S1,S2,c.match c with
250   [ mk_mconfig s t ⇒ mk_mconfig sig (FinSum S1 S2) n (inr … s) t ].
251
252 (* 
253 definition halt_liftL ≝ 
254   λS1,S2,halt.λs:FinSum S1 S2.
255   match s with
256   [ inl s1 ⇒ halt s1
257   | inr _ ⇒ true ]. (* should be vacuous in all cases we use halt_liftL *)
258
259 definition halt_liftR ≝ 
260   λS1,S2,halt.λs:FinSum S1 S2.
261   match s with
262   [ inl _ ⇒ false 
263   | inr s2 ⇒ halt s2 ]. *)
264       
265 lemma p_halt_liftL : ∀sig,n,S1,S2,halt,c.
266   halt (cstate sig S1 n c) =
267      halt_liftL S1 S2 halt (cstate … (lift_confL … c)).
268 #sig #n #S1 #S2 #halt #c cases c #s #t %
269 qed.
270
271 lemma trans_seq_liftL : ∀sig,n,M1,M2,s,a,news,move.
272   halt ?? M1 s = false → 
273   trans sig n M1 〈s,a〉 = 〈news,move〉 → 
274   trans sig n (seq sig n M1 M2) 〈inl … s,a〉 = 〈inl … news,move〉.
275 #sig #n (*#M1*) * #Q1 #T1 #init1 #halt1 #M2 #s #a #news #move
276 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
277 qed.
278
279 lemma trans_seq_liftR : ∀sig,n,M1,M2,s,a,news,move.
280   halt ?? M2 s = false → 
281   trans sig n M2 〈s,a〉 = 〈news,move〉 → 
282   trans sig n (seq sig n M1 M2) 〈inr … s,a〉 = 〈inr … news,move〉.
283 #sig #n #M1 * #Q2 #T2 #init2 #halt2 #s #a #news #move
284 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
285 qed.
286
287 lemma step_seq_liftR : ∀sig,n,M1,M2,c0.
288  halt ?? M2 (cstate ??? c0) = false → 
289  step sig n (seq sig n M1 M2) (lift_confR sig n (states ?? M1) (states ?? M2) c0) =
290  lift_confR sig n (states ?? M1) (states ?? M2) (step sig n M2 c0).
291 #sig #n #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 * #s #t
292 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
293 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
294 #s0 #m0 #Heq #Hhalt whd in ⊢ (???(?????%)); >Heq  whd in ⊢ (???%);
295 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_seq_liftR … Heq) //
296 qed.
297
298 lemma step_seq_liftL : ∀sig,n,M1,M2,c0.
299  halt ?? M1 (cstate ??? c0) = false → 
300  step sig n (seq sig n M1 M2) (lift_confL sig n (states ?? M1) (states ?? M2) c0) =
301  lift_confL sig n ?? (step sig n M1 c0).
302 #sig #n #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 * #s #t
303   lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
304   cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
305   #s0 #m0 #Heq #Hhalt
306   whd in ⊢ (???(?????%)); >Heq whd in ⊢ (???%);
307   whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_seq_liftL … Heq) //
308 qed.
309
310 lemma trans_liftL_true : ∀sig,n,M1,M2,s,a.
311   halt ?? M1 s = true → 
312   trans sig n (seq sig n M1 M2) 〈inl … s,a〉 = 〈inr … (start ?? M2),null_action sig n〉.
313 #sig #n #M1 #M2 #s #a #Hhalt whd in ⊢ (??%?); >Hhalt %
314 qed.
315
316 lemma eq_ctape_lift_conf_L : ∀sig,n,S1,S2,outc.
317   ctapes sig (FinSum S1 S2) n (lift_confL … outc) = ctapes … outc.
318 #sig #n #S1 #S2 #outc cases outc #s #t %
319 qed.
320   
321 lemma eq_ctape_lift_conf_R : ∀sig,n,S1,S2,outc.
322   ctapes sig (FinSum S1 S2) n (lift_confR … outc) = ctapes … outc.
323 #sig #n #S1 #S2 #outc cases outc #s #t %
324 qed.
325
326 theorem sem_seq: ∀sig,n.∀M1,M2:mTM sig n.∀R1,R2.
327   M1 ⊨ R1 → M2 ⊨ R2 → M1 · M2 ⊨ R1 ∘ R2.
328 #sig #n #M1 #M2 #R1 #R2 #HR1 #HR2 #t 
329 cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1
330 cases (HR2 (ctapes sig (states ?? M1) n outc1)) #k2 * #outc2 * #Hloop2 #HM2
331 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
332 %
333 [@(loop_merge ??????????? 
334    (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
335    (step sig n M1) (step sig n (seq sig n M1 M2)) 
336    (λc.halt sig n M1 (cstate … c)) 
337    (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
338   [ * *
339    [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
340    | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
341   || #c0 #Hhalt <step_seq_liftL //
342   | #x <p_halt_liftL %
343   |6:cases outc1 #s1 #t1 %
344   |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2) 
345     [ * #s2 #t2 %
346     | #c0 #Hhalt <step_seq_liftR // ]
347   |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
348    generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10 
349    >(trans_liftL_true sig n M1 M2 ??) 
350     [ whd in ⊢ (??%?); whd in ⊢ (???%);
351       @mconfig_eq whd in ⊢ (???%); // 
352     | @(loop_Some ?????? Hloop10) ]
353  ]
354 | @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
355   % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R //
356 ]
357 qed.
358
359 theorem sem_seq_app: ∀sig,n.∀M1,M2:mTM sig n.∀R1,R2,R3.
360   M1 ⊨ R1 → M2 ⊨ R2 → R1 ∘ R2 ⊆ R3 → M1 · M2 ⊨ R3.
361 #sig #n #M1 #M2 #R1 #R2 #R3 #HR1 #HR2 #Hsub
362 #t cases (sem_seq … HR1 HR2 t)
363 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
364 % [@Hloop |@Hsub @Houtc]
365 qed.
366
367 (* composition with guards *)
368 theorem sem_seq_guarded: ∀sig,n.∀M1,M2:mTM sig n.∀Pre1,Pre2,R1,R2.
369   GRealize sig n M1 Pre1 R1 → GRealize sig n M2 Pre2 R2 → 
370   (∀t1,t2.Pre1 t1 → R1 t1 t2 → Pre2 t2) → 
371   GRealize sig n (M1 · M2) Pre1 (R1 ∘ R2).
372 #sig #n #M1 #M2 #Pre1 #Pre2 #R1 #R2 #HGR1 #HGR2 #Hinv #t1 #HPre1
373 cases (HGR1 t1 HPre1) #k1 * #outc1 * #Hloop1 #HM1
374 cases (HGR2 (ctapes sig (states ?? M1) n outc1) ?) 
375   [2: @(Hinv … HPre1 HM1)]  
376 #k2 * #outc2 * #Hloop2 #HM2
377 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
378 %
379 [@(loop_merge ??????????? 
380    (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
381    (step sig n M1) (step sig n (seq sig n M1 M2)) 
382    (λc.halt sig n M1 (cstate … c)) 
383    (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
384   [ * *
385    [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
386    | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
387   || #c0 #Hhalt <step_seq_liftL //
388   | #x <p_halt_liftL %
389   |6:cases outc1 #s1 #t1 %
390   |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2) 
391     [ * #s2 #t2 %
392     | #c0 #Hhalt <step_seq_liftR // ]
393   |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
394    generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10 
395    >(trans_liftL_true sig n M1 M2 ??) 
396     [ whd in ⊢ (??%?); whd in ⊢ (???%);
397       @mconfig_eq whd in ⊢ (???%); //
398     | @(loop_Some ?????? Hloop10) ]
399  ]
400 | @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) n (lift_confL … outc1)))
401   % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R //
402 ]
403 qed.
404
405 theorem sem_seq_app_guarded: ∀sig,n.∀M1,M2:mTM sig n.∀Pre1,Pre2,R1,R2,R3.
406   GRealize sig n M1 Pre1 R1 → GRealize sig n M2 Pre2 R2 → 
407   (∀t1,t2.Pre1 t1 → R1 t1 t2 → Pre2 t2) → R1 ∘ R2 ⊆ R3 →
408   GRealize sig n (M1 · M2) Pre1 R3.
409 #sig #n #M1 #M2 #Pre1 #Pre2 #R1 #R2 #R3 #HR1 #HR2 #Hinv #Hsub
410 #t #HPre1 cases (sem_seq_guarded … HR1 HR2 Hinv t HPre1)
411 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
412 % [@Hloop |@Hsub @Houtc]
413 qed.
414
415 theorem acc_sem_seq : ∀sig,n.∀M1,M2:mTM sig n.∀R1,Rtrue,Rfalse,acc.
416   M1 ⊨ R1 → M2 ⊨ [ acc: Rtrue, Rfalse ] → 
417   M1 · M2 ⊨ [ inr … acc: R1 ∘ Rtrue, R1 ∘ Rfalse ].
418 #sig #n #M1 #M2 #R1 #Rtrue #Rfalse #acc #HR1 #HR2 #t 
419 cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1
420 cases (HR2 (ctapes sig (states ?? M1) n outc1)) #k2 * #outc2 * * #Hloop2 
421 #HMtrue #HMfalse
422 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
423 % [ %
424 [@(loop_merge ??????????? 
425    (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
426    (step sig n M1) (step sig n (seq sig n M1 M2)) 
427    (λc.halt sig n M1 (cstate … c)) 
428    (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
429   [ * *
430    [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
431    | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
432   || #c0 #Hhalt <step_seq_liftL //
433   | #x <p_halt_liftL %
434   |6:cases outc1 #s1 #t1 %
435   |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2) 
436     [ * #s2 #t2 %
437     | #c0 #Hhalt <step_seq_liftR // ]
438   |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
439    generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10 
440    >(trans_liftL_true sig n M1 M2 ??) 
441     [ whd in ⊢ (??%?); whd in ⊢ (???%);
442       @mconfig_eq whd in ⊢ (???%); // 
443     | @(loop_Some ?????? Hloop10) ]
444  ]
445 | >(mconfig_expand … outc2) in ⊢ (%→?); whd in ⊢ (??%?→?); 
446   #Hqtrue destruct (Hqtrue)
447   @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
448   % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R /2/ ]
449 | >(mconfig_expand … outc2) in ⊢ (%→?); whd in ⊢ (?(??%?)→?); #Hqfalse
450   @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
451   % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R @HMfalse
452   @(not_to_not … Hqfalse) //
453 ]
454 qed.
455
456 lemma acc_sem_seq_app : ∀sig,n.∀M1,M2:mTM sig n.∀R1,Rtrue,Rfalse,R2,R3,acc.
457   M1 ⊨ R1 → M2 ⊨ [acc: Rtrue, Rfalse] → 
458     (∀t1,t2,t3. R1 t1 t3 → Rtrue t3 t2 → R2 t1 t2) → 
459     (∀t1,t2,t3. R1 t1 t3 → Rfalse t3 t2 → R3 t1 t2) → 
460     M1 · M2 ⊨ [inr … acc : R2, R3].    
461 #sig #n #M1 #M2 #R1 #Rtrue #Rfalse #R2 #R3 #acc
462 #HR1 #HRacc #Hsub1 #Hsub2 
463 #t cases (acc_sem_seq … HR1 HRacc t)
464 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
465 % [% [@Hloop
466      |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
467   |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]
468 qed.