1 include "turing/mono.ma".
2 include "basics/vectors.ma".
4 (* We do not distinuish an input tape *)
6 (* tapes_no = number of ADDITIONAL working tapes *)
8 record mTM (sig:FinSet) (tapes_no:nat) : Type[1] ≝
10 trans : states × (Vector (option sig) (S tapes_no)) →
11 states × (Vector ((option sig) × move) (S tapes_no));
16 record mconfig (sig,states:FinSet) (n:nat): Type[0] ≝
18 ctapes : Vector (tape sig) (S n)
21 lemma mconfig_expand: ∀sig,n,Q,c.
22 c = mk_mconfig sig Q n (cstate ??? c) (ctapes ??? c).
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 //
32 definition current_chars ≝ λsig.λn.λtapes.
33 vec_map ?? (current sig) (S n) tapes.
35 lemma nth_current_chars : ∀sig,n,tapes,i.
36 nth i ? (current_chars sig n tapes) (None ?)
37 = current sig (nth i ? tapes (niltape sig)).
38 #sig #n #tapes #i >(nth_vec_map … (current sig) i (S n)) %
41 definition tape_move_multi ≝
43 pmap_vec ??? (tape_move_mono sig) n ts mvs.
45 lemma tape_move_multi_def : ∀sig,n,ts,mvs.
46 tape_move_multi sig n ts mvs = pmap_vec ??? (tape_move_mono sig) n ts mvs.
49 definition step ≝ λsig.λn.λM:mTM sig n.λc:mconfig sig (states ?? M) n.
50 let 〈news,mvs〉 ≝ trans sig n M 〈cstate ??? c,current_chars ?? (ctapes ??? c)〉 in
51 mk_mconfig ??? news (tape_move_multi sig ? (ctapes ??? c) mvs).
53 definition empty_tapes ≝ λsig.λn.
54 mk_Vector ? n (make_list (tape sig) (niltape sig) n) ?.
55 elim n // normalize //
58 (************************** Realizability *************************************)
59 definition loopM ≝ λsig,n.λM:mTM sig n.λi,cin.
60 loop ? i (step sig n M) (λc.halt sig n M (cstate ??? c)) cin.
62 lemma loopM_unfold : ∀sig,n,M,i,cin.
63 loopM sig n M i cin = loop ? i (step sig n M) (λc.halt sig n M (cstate ??? c)) cin.
66 definition initc ≝ λsig,n.λM:mTM sig n.λtapes.
67 mk_mconfig sig (states sig n M) n (start sig n M) tapes.
69 definition Realize ≝ λsig,n.λM:mTM sig n.λR:relation (Vector (tape sig) ?).
71 loopM sig n M i (initc sig n M t) = Some ? outc ∧ R t (ctapes ??? outc).
73 definition WRealize ≝ λsig,n.λM:mTM sig n.λR:relation (Vector (tape sig) ?).
75 loopM sig n M i (initc sig n M t) = Some ? outc → R t (ctapes ??? outc).
77 definition Terminate ≝ λsig,n.λM:mTM sig n.λt. ∃i,outc.
78 loopM sig n M i (initc sig n M t) = Some ? outc.
80 (* notation "M \vDash R" non associative with precedence 45 for @{ 'models $M $R}. *)
81 interpretation "multi realizability" 'models M R = (Realize ?? M R).
83 (* notation "M \VDash R" non associative with precedence 45 for @{ 'wmodels $M $R}. *)
84 interpretation "weak multi realizability" 'wmodels M R = (WRealize ?? M R).
86 interpretation "multi termination" 'fintersects M t = (Terminate ?? M t).
88 lemma WRealize_to_Realize : ∀sig,n .∀M: mTM sig n.∀R.
89 (∀t.M ↓ t) → M ⊫ R → M ⊨ R.
90 #sig #n #M #R #HT #HW #t cases (HT … t) #i * #outc #Hloop
91 @(ex_intro … i) @(ex_intro … outc) % // @(HW … i) //
94 theorem Realize_to_WRealize : ∀sig,n.∀M:mTM sig n.∀R.
96 #sig #n #M #R #H1 #inc #i #outc #Hloop
97 cases (H1 inc) #k * #outc1 * #Hloop1 #HR >(loop_eq … Hloop Hloop1) //
100 definition accRealize ≝ λsig,n.λM:mTM sig n.λacc:states sig n M.λRtrue,Rfalse.
102 loopM sig n M i (initc sig n M t) = Some ? outc ∧
103 (cstate ??? outc = acc → Rtrue t (ctapes ??? outc)) ∧
104 (cstate ??? outc ≠ acc → Rfalse t (ctapes ??? outc)).
106 (* notation "M ⊨ [q: R1,R2]" non associative with precedence 45 for @{ 'cmodels $M $q $R1 $R2}. *)
107 interpretation "conditional multi realizability" 'cmodels M q R1 R2 = (accRealize ?? M q R1 R2).
109 (*************************** guarded realizablity *****************************)
110 definition GRealize ≝ λsig,n.λM:mTM sig n.
111 λPre:Vector (tape sig) ? →Prop.λR:relation (Vector (tape sig) ?).
113 loopM sig n M i (initc sig n M t) = Some ? outc ∧ R t (ctapes ??? outc).
115 definition accGRealize ≝ λsig,n.λM:mTM sig n.λacc:states sig n M.
116 λPre: Vector (tape sig) ? → Prop.λRtrue,Rfalse.
118 loopM sig n M i (initc sig n M t) = Some ? outc ∧
119 (cstate ??? outc = acc → Rtrue t (ctapes ??? outc)) ∧
120 (cstate ??? outc ≠ acc → Rfalse t (ctapes ??? outc)).
122 lemma WRealize_to_GRealize : ∀sig,n.∀M: mTM sig n.∀Pre,R.
123 (∀t.Pre t → M ↓ t) → M ⊫ R → GRealize sig n M Pre R.
124 #sig #n #M #Pre #R #HT #HW #t #HPre cases (HT … t HPre) #i * #outc #Hloop
125 @(ex_intro … i) @(ex_intro … outc) % // @(HW … i) //
128 lemma Realize_to_GRealize : ∀sig,n.∀M: mTM sig n.∀P,R.
129 M ⊨ R → GRealize sig n M P R.
130 #alpha #n #M #Pre #R #HR #t #HPre
131 cases (HR t) -HR #k * #outc * #Hloop #HR
132 @(ex_intro ?? k) @(ex_intro ?? outc) %
136 lemma acc_Realize_to_acc_GRealize: ∀sig,n.∀M:mTM sig n.∀q:states sig n M.∀P,R1,R2.
137 M ⊨ [q:R1,R2] → accGRealize sig n M q P R1 R2.
138 #alpha #n #M #q #Pre #R1 #R2 #HR #t #HPre
139 cases (HR t) -HR #k * #outc * * #Hloop #HRtrue #HRfalse
140 @(ex_intro ?? k) @(ex_intro ?? outc) %
141 [ % [@Hloop] @HRtrue | @HRfalse]
144 (******************************** monotonicity ********************************)
145 lemma Realize_to_Realize : ∀sig,n.∀M:mTM sig n.∀R1,R2.
146 R1 ⊆ R2 → M ⊨ R1 → M ⊨ R2.
147 #alpha #n #M #R1 #R2 #Himpl #HR1 #intape
148 cases (HR1 intape) -HR1 #k * #outc * #Hloop #HR1
149 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
152 lemma WRealize_to_WRealize: ∀sig,n.∀M:mTM sig n.∀R1,R2.
153 R1 ⊆ R2 → WRealize sig n M R1 → WRealize sig n M R2.
154 #alpha #n #M #R1 #R2 #Hsub #HR1 #intape #i #outc #Hloop
155 @Hsub @(HR1 … i) @Hloop
158 lemma GRealize_to_GRealize : ∀sig,n.∀M:mTM sig n.∀P,R1,R2.
159 R1 ⊆ R2 → GRealize sig n M P R1 → GRealize sig n M P R2.
160 #alpha #n #M #P #R1 #R2 #Himpl #HR1 #intape #HP
161 cases (HR1 intape HP) -HR1 #k * #outc * #Hloop #HR1
162 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
165 lemma GRealize_to_GRealize_2 : ∀sig,n.∀M:mTM sig n.∀P1,P2,R1,R2.
166 P2 ⊆ P1 → R1 ⊆ R2 → GRealize sig n M P1 R1 → GRealize sig n M P2 R2.
167 #alpha #n #M #P1 #P2 #R1 #R2 #Himpl1 #Himpl2 #H1 #intape #HP
168 cases (H1 intape (Himpl1 … HP)) -H1 #k * #outc * #Hloop #H1
169 @(ex_intro ?? k) @(ex_intro ?? outc) % /2/
172 lemma acc_Realize_to_acc_Realize: ∀sig,n.∀M:mTM sig n.∀q:states sig n M.
174 R1 ⊆ R3 → R2 ⊆ R4 → M ⊨ [q:R1,R2] → M ⊨ [q:R3,R4].
175 #alpha #n #M #q #R1 #R2 #R3 #R4 #Hsub13 #Hsub24 #HRa #intape
176 cases (HRa intape) -HRa #k * #outc * * #Hloop #HRtrue #HRfalse
177 @(ex_intro ?? k) @(ex_intro ?? outc) %
178 [ % [@Hloop] #Hq @Hsub13 @HRtrue // | #Hq @Hsub24 @HRfalse //]
181 (**************************** A canonical relation ****************************)
183 definition R_mTM ≝ λsig,n.λM:mTM sig n.λq.λt1,t2.
185 loopM ? n M i (mk_mconfig ??? q t1) = Some ? outc ∧
186 t2 = (ctapes ??? outc).
188 lemma R_mTM_to_R: ∀sig,n.∀M:mTM sig n.∀R. ∀t1,t2.
189 M ⊫ R → R_mTM ?? M (start sig n M) t1 t2 → R t1 t2.
190 #sig #n #M #R #t1 #t2 whd in ⊢ (%→?); #HMR * #i * #outc *
191 #Hloop #Ht2 >Ht2 @(HMR … Hloop)
194 (******************************** NOP Machine *********************************)
199 definition nop_states ≝ initN 1.
200 definition start_nop : initN 1 ≝ mk_Sig ?? 0 (le_n … 1). *)
203 λalpha:FinSet.λn.mk_mTM alpha n nop_states
204 (λp.let 〈q,a〉 ≝ p in 〈q,mk_Vector ? (S n) (make_list ? (〈None ?,N〉) (S n)) ?〉)
209 definition R_nop ≝ λalpha,n.λt1,t2:Vector (tape alpha) (S n).t2 = t1.
212 ∀alpha,n.nop alpha n⊨ R_nop alpha n.
213 #alpha #n #intapes @(ex_intro ?? 1)
214 @(ex_intro … (mk_mconfig ??? start_nop intapes)) % %
217 lemma nop_single_state: ∀sig,n.∀q1,q2:states ? n (nop sig n). q1 = q2.
218 normalize #sig #n0 * #n #ltn1 * #m #ltm1
219 generalize in match ltn1; generalize in match ltm1;
220 <(le_n_O_to_eq … (le_S_S_to_le … ltn1)) <(le_n_O_to_eq … (le_S_S_to_le … ltm1))
223 (************************** Sequential Composition ****************************)
224 definition null_action ≝ λsig.λn.
225 mk_Vector ? (S n) (make_list (option sig × move) (〈None ?,N〉) (S n)) ?.
226 elim (S n) // normalize //
229 lemma tape_move_null_action: ∀sig,n,tapes.
230 tape_move_multi sig (S n) tapes (null_action sig n) = tapes.
231 #sig #n #tapes cases tapes -tapes #tapes whd in match (null_action ??);
232 #Heq @Vector_eq <Heq -Heq elim tapes //
233 #a #tl #Hind whd in ⊢ (??%?); @eq_f2 // @Hind
236 definition seq_trans ≝ λsig,n. λM1,M2 : mTM sig n.
240 if halt sig n M1 s1 then 〈inr … (start sig n M2), null_action sig n〉
241 else let 〈news1,m〉 ≝ trans sig n M1 〈s1,a〉 in 〈inl … news1,m〉
242 | inr s2 ⇒ let 〈news2,m〉 ≝ trans sig n M2 〈s2,a〉 in 〈inr … news2,m〉
245 definition seq ≝ λsig,n. λM1,M2 : mTM sig n.
247 (FinSum (states sig n M1) (states sig n M2))
248 (seq_trans sig n M1 M2)
249 (inl … (start sig n M1))
251 [ inl _ ⇒ false | inr s2 ⇒ halt sig n M2 s2]).
253 (* notation "a · b" right associative with precedence 65 for @{ 'middot $a $b}. *)
254 interpretation "sequential composition" 'middot a b = (seq ?? a b).
256 definition lift_confL ≝
257 λsig,n,S1,S2,c.match c with
258 [ mk_mconfig s t ⇒ mk_mconfig sig (FinSum S1 S2) n (inl … s) t ].
260 definition lift_confR ≝
261 λsig,n,S1,S2,c.match c with
262 [ mk_mconfig s t ⇒ mk_mconfig sig (FinSum S1 S2) n (inr … s) t ].
265 definition halt_liftL ≝
266 λS1,S2,halt.λs:FinSum S1 S2.
269 | inr _ ⇒ true ]. (* should be vacuous in all cases we use halt_liftL *)
271 definition halt_liftR ≝
272 λS1,S2,halt.λs:FinSum S1 S2.
275 | inr s2 ⇒ halt s2 ]. *)
277 lemma p_halt_liftL : ∀sig,n,S1,S2,halt,c.
278 halt (cstate sig S1 n c) =
279 halt_liftL S1 S2 halt (cstate … (lift_confL … c)).
280 #sig #n #S1 #S2 #halt #c cases c #s #t %
283 lemma trans_seq_liftL : ∀sig,n,M1,M2,s,a,news,move.
284 halt ?? M1 s = false →
285 trans sig n M1 〈s,a〉 = 〈news,move〉 →
286 trans sig n (seq sig n M1 M2) 〈inl … s,a〉 = 〈inl … news,move〉.
287 #sig #n (*#M1*) * #Q1 #T1 #init1 #halt1 #M2 #s #a #news #move
288 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
291 lemma trans_seq_liftR : ∀sig,n,M1,M2,s,a,news,move.
292 halt ?? M2 s = false →
293 trans sig n M2 〈s,a〉 = 〈news,move〉 →
294 trans sig n (seq sig n M1 M2) 〈inr … s,a〉 = 〈inr … news,move〉.
295 #sig #n #M1 * #Q2 #T2 #init2 #halt2 #s #a #news #move
296 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
299 lemma step_seq_liftR : ∀sig,n,M1,M2,c0.
300 halt ?? M2 (cstate ??? c0) = false →
301 step sig n (seq sig n M1 M2) (lift_confR sig n (states ?? M1) (states ?? M2) c0) =
302 lift_confR sig n (states ?? M1) (states ?? M2) (step sig n M2 c0).
303 #sig #n #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 * #s #t
304 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
305 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
306 #s0 #m0 #Heq #Hhalt whd in ⊢ (???(?????%)); >Heq whd in ⊢ (???%);
307 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_seq_liftR … Heq) //
310 lemma step_seq_liftL : ∀sig,n,M1,M2,c0.
311 halt ?? M1 (cstate ??? c0) = false →
312 step sig n (seq sig n M1 M2) (lift_confL sig n (states ?? M1) (states ?? M2) c0) =
313 lift_confL sig n ?? (step sig n M1 c0).
314 #sig #n #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 * #s #t
315 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
316 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
318 whd in ⊢ (???(?????%)); >Heq whd in ⊢ (???%);
319 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_seq_liftL … Heq) //
322 lemma trans_liftL_true : ∀sig,n,M1,M2,s,a.
323 halt ?? M1 s = true →
324 trans sig n (seq sig n M1 M2) 〈inl … s,a〉 = 〈inr … (start ?? M2),null_action sig n〉.
325 #sig #n #M1 #M2 #s #a #Hhalt whd in ⊢ (??%?); >Hhalt %
328 lemma eq_ctape_lift_conf_L : ∀sig,n,S1,S2,outc.
329 ctapes sig (FinSum S1 S2) n (lift_confL … outc) = ctapes … outc.
330 #sig #n #S1 #S2 #outc cases outc #s #t %
333 lemma eq_ctape_lift_conf_R : ∀sig,n,S1,S2,outc.
334 ctapes sig (FinSum S1 S2) n (lift_confR … outc) = ctapes … outc.
335 #sig #n #S1 #S2 #outc cases outc #s #t %
338 theorem sem_seq: ∀sig,n.∀M1,M2:mTM sig n.∀R1,R2.
339 M1 ⊨ R1 → M2 ⊨ R2 → M1 · M2 ⊨ R1 ∘ R2.
340 #sig #n #M1 #M2 #R1 #R2 #HR1 #HR2 #t
341 cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1
342 cases (HR2 (ctapes sig (states ?? M1) n outc1)) #k2 * #outc2 * #Hloop2 #HM2
343 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
345 [@(loop_merge ???????????
346 (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
347 (step sig n M1) (step sig n (seq sig n M1 M2))
348 (λc.halt sig n M1 (cstate … c))
349 (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
351 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
352 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
353 || #c0 #Hhalt <step_seq_liftL //
355 |6:cases outc1 #s1 #t1 %
356 |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2)
358 | #c0 #Hhalt <step_seq_liftR // ]
359 |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
360 generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10
361 >(trans_liftL_true sig n M1 M2 ??)
362 [ whd in ⊢ (??%?); whd in ⊢ (???%);
363 @mconfig_eq whd in ⊢ (???%); //
364 | @(loop_Some ?????? Hloop10) ]
366 | @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
367 % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R //
371 theorem sem_seq_app: ∀sig,n.∀M1,M2:mTM sig n.∀R1,R2,R3.
372 M1 ⊨ R1 → M2 ⊨ R2 → R1 ∘ R2 ⊆ R3 → M1 · M2 ⊨ R3.
373 #sig #n #M1 #M2 #R1 #R2 #R3 #HR1 #HR2 #Hsub
374 #t cases (sem_seq … HR1 HR2 t)
375 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
376 % [@Hloop |@Hsub @Houtc]
379 (* composition with guards *)
380 theorem sem_seq_guarded: ∀sig,n.∀M1,M2:mTM sig n.∀Pre1,Pre2,R1,R2.
381 GRealize sig n M1 Pre1 R1 → GRealize sig n M2 Pre2 R2 →
382 (∀t1,t2.Pre1 t1 → R1 t1 t2 → Pre2 t2) →
383 GRealize sig n (M1 · M2) Pre1 (R1 ∘ R2).
384 #sig #n #M1 #M2 #Pre1 #Pre2 #R1 #R2 #HGR1 #HGR2 #Hinv #t1 #HPre1
385 cases (HGR1 t1 HPre1) #k1 * #outc1 * #Hloop1 #HM1
386 cases (HGR2 (ctapes sig (states ?? M1) n outc1) ?)
387 [2: @(Hinv … HPre1 HM1)]
388 #k2 * #outc2 * #Hloop2 #HM2
389 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
391 [@(loop_merge ???????????
392 (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
393 (step sig n M1) (step sig n (seq sig n M1 M2))
394 (λc.halt sig n M1 (cstate … c))
395 (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
397 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
398 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
399 || #c0 #Hhalt <step_seq_liftL //
401 |6:cases outc1 #s1 #t1 %
402 |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2)
404 | #c0 #Hhalt <step_seq_liftR // ]
405 |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
406 generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10
407 >(trans_liftL_true sig n M1 M2 ??)
408 [ whd in ⊢ (??%?); whd in ⊢ (???%);
409 @mconfig_eq whd in ⊢ (???%); //
410 | @(loop_Some ?????? Hloop10) ]
412 | @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) n (lift_confL … outc1)))
413 % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R //
417 theorem sem_seq_app_guarded: ∀sig,n.∀M1,M2:mTM sig n.∀Pre1,Pre2,R1,R2,R3.
418 GRealize sig n M1 Pre1 R1 → GRealize sig n M2 Pre2 R2 →
419 (∀t1,t2.Pre1 t1 → R1 t1 t2 → Pre2 t2) → R1 ∘ R2 ⊆ R3 →
420 GRealize sig n (M1 · M2) Pre1 R3.
421 #sig #n #M1 #M2 #Pre1 #Pre2 #R1 #R2 #R3 #HR1 #HR2 #Hinv #Hsub
422 #t #HPre1 cases (sem_seq_guarded … HR1 HR2 Hinv t HPre1)
423 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
424 % [@Hloop |@Hsub @Houtc]
427 theorem acc_sem_seq : ∀sig,n.∀M1,M2:mTM sig n.∀R1,Rtrue,Rfalse,acc.
428 M1 ⊨ R1 → M2 ⊨ [ acc: Rtrue, Rfalse ] →
429 M1 · M2 ⊨ [ inr … acc: R1 ∘ Rtrue, R1 ∘ Rfalse ].
430 #sig #n #M1 #M2 #R1 #Rtrue #Rfalse #acc #HR1 #HR2 #t
431 cases (HR1 t) #k1 * #outc1 * #Hloop1 #HM1
432 cases (HR2 (ctapes sig (states ?? M1) n outc1)) #k2 * #outc2 * * #Hloop2
434 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … outc2))
436 [@(loop_merge ???????????
437 (loop_lift ??? (lift_confL sig n (states sig n M1) (states sig n M2))
438 (step sig n M1) (step sig n (seq sig n M1 M2))
439 (λc.halt sig n M1 (cstate … c))
440 (λc.halt_liftL ?? (halt sig n M1) (cstate … c)) … Hloop1))
442 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
443 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
444 || #c0 #Hhalt <step_seq_liftL //
446 |6:cases outc1 #s1 #t1 %
447 |7:@(loop_lift … (initc ??? (ctapes … outc1)) … Hloop2)
449 | #c0 #Hhalt <step_seq_liftR // ]
450 |whd in ⊢ (??(????%)?);whd in ⊢ (??%?);
451 generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10
452 >(trans_liftL_true sig n M1 M2 ??)
453 [ whd in ⊢ (??%?); whd in ⊢ (???%);
454 @mconfig_eq whd in ⊢ (???%); //
455 | @(loop_Some ?????? Hloop10) ]
457 | >(mconfig_expand … outc2) in ⊢ (%→?); whd in ⊢ (??%?→?);
458 #Hqtrue destruct (Hqtrue)
459 @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
460 % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R /2/ ]
461 | >(mconfig_expand … outc2) in ⊢ (%→?); whd in ⊢ (?(??%?)→?); #Hqfalse
462 @(ex_intro … (ctapes ? (FinSum (states ?? M1) (states ?? M2)) ? (lift_confL … outc1)))
463 % // >eq_ctape_lift_conf_L >eq_ctape_lift_conf_R @HMfalse
464 @(not_to_not … Hqfalse) //
468 lemma acc_sem_seq_app : ∀sig,n.∀M1,M2:mTM sig n.∀R1,Rtrue,Rfalse,R2,R3,acc.
469 M1 ⊨ R1 → M2 ⊨ [acc: Rtrue, Rfalse] →
470 (∀t1,t2,t3. R1 t1 t3 → Rtrue t3 t2 → R2 t1 t2) →
471 (∀t1,t2,t3. R1 t1 t3 → Rfalse t3 t2 → R3 t1 t2) →
472 M1 · M2 ⊨ [inr … acc : R2, R3].
473 #sig #n #M1 #M2 #R1 #Rtrue #Rfalse #R2 #R3 #acc
474 #HR1 #HRacc #Hsub1 #Hsub2
475 #t cases (acc_sem_seq … HR1 HRacc t)
476 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
478 |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
479 |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]