2 ||M|| This file is part of HELM, an Hypertextual, Electronic
3 ||A|| Library of Mathematics, developed at the Computer Science
4 ||T|| Department of the University of Bologna, Italy.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_____________________________________________________________*)
12 include "turing/turing.ma".
14 (**************************** single final machine ****************************)
16 definition single_finalTM ≝
17 λsig,n.λM:mTM sig n.seq ?? M (nop ??).
19 lemma sem_single_final: ∀sig,n.∀M: mTM sig n.∀R.
20 M ⊨ R → single_finalTM sig n M ⊨ R.
21 #sig #n #M #R #HR #intape
22 cases (sem_seq ?????? HR (sem_nop …) intape)
23 #k * #outc * #Hloop * #ta * #Hta whd in ⊢ (%→?); #Houtc
24 @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop | >Houtc // ]
27 lemma single_final: ∀sig,n.∀M: mTM sig n.∀q1,q2.
28 halt ?? (single_finalTM sig n M) q1 = true
29 → halt ?? (single_finalTM sig n M) q2 = true → q1=q2.
31 [#q1M #q2 whd in match (halt ????); #H destruct
33 [#q2M #_ whd in match (halt ????); #H destruct
34 |#q2nop #_ #_ @eq_f normalize @nop_single_state
39 (******************************** if machine **********************************)
41 definition if_trans ≝ λsig,n. λM1,M2,M3 : mTM sig n. λq:states sig n M1.
45 if halt sig n M1 s1 then
46 if s1==q then 〈inr … (inl … (start sig n M2)), null_action ??〉
47 else 〈inr … (inr … (start sig n M3)), null_action ??〉
48 else let 〈news1,m〉 ≝ trans sig n M1 〈s1,a〉 in
52 [ inl s2 ⇒ let 〈news2,m〉 ≝ trans sig n M2 〈s2,a〉 in
53 〈inr … (inl … news2),m〉
54 | inr s3 ⇒ let 〈news3,m〉 ≝ trans sig n M3 〈s3,a〉 in
55 〈inr … (inr … news3),m〉
59 definition ifTM ≝ λsig,n. λcondM,thenM,elseM : mTM sig n.
60 λqacc: states sig n condM.
62 (FinSum (states sig n condM) (FinSum (states sig n thenM) (states sig n elseM)))
63 (if_trans sig n condM thenM elseM qacc)
64 (inl … (start sig n condM))
67 | inr s' ⇒ match s' with
68 [ inl s2 ⇒ halt sig n thenM s2
69 | inr s3 ⇒ halt sig n elseM s3 ]]).
71 (****************************** lifting lemmas ********************************)
72 lemma trans_if_liftM1 : ∀sig,n,M1,M2,M3,acc,s,a,news,move.
73 halt ?? M1 s = false →
74 trans sig n M1 〈s,a〉 = 〈news,move〉 →
75 trans sig n (ifTM sig n M1 M2 M3 acc) 〈inl … s,a〉 = 〈inl … news,move〉.
76 #sig #n * #Q1 #T1 #init1 #halt1 #M2 #M3 #acc #s #a #news #move
77 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
80 lemma trans_if_liftM2 : ∀sig,n,M1,M2,M3,acc,s,a,news,move.
81 halt ?? M2 s = false →
82 trans sig n M2 〈s,a〉 = 〈news,move〉 →
83 trans sig n (ifTM sig n M1 M2 M3 acc) 〈inr … (inl … s),a〉 = 〈inr… (inl … news),move〉.
84 #sig #n #M1 * #Q2 #T2 #init2 #halt2 #M3 #acc #s #a #news #move
85 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
88 lemma trans_if_liftM3 : ∀sig,n,M1,M2,M3,acc,s,a,news,move.
89 halt ?? M3 s = false →
90 trans sig n M3 〈s,a〉 = 〈news,move〉 →
91 trans sig n (ifTM sig n M1 M2 M3 acc) 〈inr … (inr … s),a〉 = 〈inr… (inr … news),move〉.
92 #sig #n #M1 * #Q2 #T2 #init2 #halt2 #M3 #acc #s #a #news #move
93 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
96 lemma step_if_liftM1 : ∀sig,n,M1,M2,M3,acc,c0.
97 halt ?? M1 (cstate ??? c0) = false →
98 step sig n (ifTM sig n M1 M2 M3 acc) (lift_confL sig n (states ?? M1) ? c0) =
99 lift_confL sig n (states ?? M1) ? (step sig n M1 c0).
100 #sig #n #M1 #M2 #M3 #acc * #s #t
101 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
102 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
104 whd in ⊢ (???(?????%)); >Heq whd in ⊢ (???%);
105 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_if_liftM1 … Hhalt Heq) //
108 lemma step_if_liftM2 : ∀sig,n,M1,M2,M3,acc,c0.
109 halt ?? M2 (cstate ??? c0) = false →
110 step sig n (ifTM sig n M1 M2 M3 acc) (lift_confR sig ??? (lift_confL sig ??? c0)) =
111 lift_confR sig ??? (lift_confL sig ??? (step sig n M2 c0)).
112 #sig #n #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 #M3 #acc * #s #t
113 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
114 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
116 whd in match (step ?? M2 ?); >Heq whd in ⊢ (???%);
117 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_if_liftM2 … Hhalt Heq) //
120 lemma step_if_liftM3 : ∀sig,n,M1,M2,M3,acc,c0.
121 halt ?? M3 (cstate ??? c0) = false →
122 step sig n (ifTM sig n M1 M2 M3 acc) (lift_confR sig ??? (lift_confR sig ??? c0)) =
123 lift_confR sig ??? (lift_confR sig ??? (step sig n M3 c0)).
124 #sig #n #M1 #M2 #M3 #acc * #s #t
125 lapply (refl ? (trans ??? 〈s,current_chars sig n t〉))
126 cases (trans ??? 〈s,current_chars sig n t〉) in ⊢ (???% → %);
128 whd in match (step ?? M3 ?); >Heq whd in ⊢ (???%);
129 whd in ⊢ (??(????%)?); whd in ⊢ (??%?); >(trans_if_liftM3 … Hhalt Heq) //
132 lemma trans_if_M1true_acc : ∀sig,n,M1,M2,M3,acc,s,a.
133 halt ?? M1 s = true → s==acc = true →
134 trans sig n (ifTM sig n M1 M2 M3 acc) 〈inl … s,a〉 =
135 〈inr … (inl … (start ?? M2)),null_action ??〉.
136 #sig #n #M1 #M2 #M3 #acc #s #a #Hhalt #Hacc whd in ⊢ (??%?); >Hhalt >Hacc %
139 lemma trans_if_M1true_notacc : ∀sig,n,M1,M2,M3,acc,s,a.
140 halt ?? M1 s = true → s==acc = false →
141 trans sig n (ifTM sig n M1 M2 M3 acc) 〈inl … s,a〉 =
142 〈inr … (inr … (start ?? M3)),null_action ??〉.
143 #sig #n #M1 #M2 #M3 #acc #s #a #Hhalt #Hacc whd in ⊢ (??%?); >Hhalt >Hacc %
146 (******************************** semantics ***********************************)
147 lemma sem_if: ∀sig,n.∀M1,M2,M3:mTM sig n.∀Rtrue,Rfalse,R2,R3,acc.
148 M1 ⊨ [acc: Rtrue,Rfalse] → M2 ⊨ R2 → M3 ⊨ R3 →
149 ifTM sig n M1 M2 M3 acc ⊨ (Rtrue ∘ R2) ∪ (Rfalse ∘ R3).
150 #sig #n #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HR3 #t
151 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse
152 cases (true_or_false (cstate ??? outc1 == acc)) #Hacc
153 [cases (HR2 (ctapes sig ?? outc1)) #k2 * #outc2 * #Hloop2 #HM2
154 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confL … outc2))) %
155 [@(loop_merge ?????????
156 (mk_mconfig ? (FinSum (states sig n M1) (FinSum (states sig n M2) (states sig n M3))) n
157 (inr (states sig n M1) ? (inl (states sig n M2) (states sig n M3) (start sig n M2))) (ctapes ??? outc1) )
160 (lift_confL sig n (states ?? M1) (FinSum (states ?? M2) (states ?? M3)))
161 (step sig n M1) (step sig n (ifTM sig n M1 M2 M3 acc))
162 (λc.halt sig n M1 (cstate … c))
163 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
166 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
167 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
168 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
170 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1)
171 whd in match (lift_confL ?????);
172 >(trans_if_M1true_acc … Hacc)
173 [@mconfig_eq // (* whd in ⊢ (??%?); *)
174 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
175 |@(loop_Some ?????? Hloop1)]
176 |cases outc1 #s1 #t1 %
178 (λc.(lift_confR … (lift_confL sig n (states ?? M2) (states ?? M3) c)))
181 | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
183 |%1 @(ex_intro … (ctapes ??? outc1)) %
184 [@HMtrue @(\P Hacc) | >(mconfig_expand ??? outc2) @HM2 ]
186 |cases (HR3 (ctapes sig ?? outc1)) #k2 * #outc2 * #Hloop2 #HM3
187 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
188 [@(loop_merge ?????????
189 (mk_mconfig ? (FinSum (states sig ? M1) (FinSum (states sig ? M2) (states sig ? M3))) n
190 (inr (states sig ? M1) ? (inr (states sig ? M2) (states sig ? M3) (start sig ? M3))) (ctapes ??? outc1) )
193 (lift_confL sig n (states ?? M1) (FinSum (states ?? M2) (states ?? M3)))
194 (step sig n M1) (step sig n (ifTM sig n M1 M2 M3 acc))
195 (λc.halt sig n M1 (cstate … c))
196 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
199 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
200 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
201 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
203 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1);
204 whd in match (lift_confL ?????);
205 >(trans_if_M1true_notacc … Hacc)
206 [@mconfig_eq // (* whd in ⊢ (??%?); *)
207 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
208 |@(loop_Some ?????? Hloop1)]
209 |cases outc1 #s1 #t1 %
211 (λc.(lift_confR … (lift_confR sig n (states ?? M2) (states ?? M3) c)))
214 | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
216 |%2 @(ex_intro … (ctapes ??? outc1)) %
217 [@HMfalse @(\Pf Hacc) | >(mconfig_expand ??? outc2) @HM3 ]
222 lemma sem_if_app: ∀sig,n,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,acc.
223 accRealize sig n M1 acc Rtrue Rfalse → M2 ⊨ R2 → M3 ⊨ R3 →
224 (∀t1,t2,t3. (Rtrue t1 t3 → R2 t3 t2) ∨ (Rfalse t1 t3 → R3 t3 t2) → R4 t1 t2) →
225 ifTM sig n M1 M2 M3 acc ⊨ R4.
226 #sig #n #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #acc
227 #HRacc #HRtrue #HRfalse #Hsub
228 #t cases (sem_if … HRacc HRtrue HRfalse t)
229 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
230 % [@Hloop] cases Houtc
231 [* #t3 * #Hleft #Hright @(Hsub … t3) %1 /2/
232 |* #t3 * #Hleft #Hright @(Hsub … t3) %2 /2/ ]
235 (* we can probably use acc_sem_if to prove sem_if *)
236 (* for sure we can use acc_sem_if_guarded to prove acc_sem_if *)
237 lemma acc_sem_if: ∀sig,n,M1,M2,M3,Rtrue,Rfalse,R2,R3,acc.
238 M1 ⊨ [acc: Rtrue, Rfalse] → M2 ⊨ R2 → M3 ⊨ R3 →
239 ifTM sig n M1 (single_finalTM … M2) M3 acc ⊨
240 [inr … (inl … (inr … start_nop)): Rtrue ∘ R2, Rfalse ∘ R3].
241 #sig #n #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HR3 #t
242 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse
243 cases (true_or_false (cstate ??? outc1 == acc)) #Hacc
244 [lapply (sem_single_final … HR2) -HR2 #HR2
245 cases (HR2 (ctapes sig ?? outc1)) #k2 * #outc2 * #Hloop2 #HM2
246 @(ex_intro … (k1+k2))
247 @(ex_intro … (lift_confR … (lift_confL … outc2))) %
249 [@(loop_merge ?????????
250 (mk_mconfig ? (states sig n (ifTM sig n M1 (single_finalTM … M2) M3 acc)) n
251 (inr (states sig n M1) ? (inl ? (states sig n M3) (start sig n (single_finalTM sig n M2)))) (ctapes ??? outc1) )
254 (lift_confL sig n (states ?? M1) (FinSum (states ?? (single_finalTM … M2)) (states ?? M3)))
255 (step sig n M1) (step sig n (ifTM sig n M1 (single_finalTM ?? M2) M3 acc))
256 (λc.halt sig n M1 (cstate … c))
257 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
260 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
261 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
262 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
264 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1);
265 whd in match (lift_confL ????);
266 >(trans_if_M1true_acc … Hacc)
267 [@mconfig_eq // (* whd in ⊢ (??%?); *)
268 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
269 | @(loop_Some ?????? Hloop1)]
270 |cases outc1 #s1 #t1 %
272 (λc.(lift_confR … (lift_confL sig n (states ?? (single_finalTM ?? M2)) (states ?? M3) c)))
275 | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
277 |#_ @(ex_intro … (ctapes ??? outc1)) %
278 [@HMtrue @(\P Hacc) | >(mconfig_expand ??? outc2) @HM2 ]
280 |>(mconfig_expand ??? outc2) whd in match (lift_confR ?????);
281 * #H @False_ind @H @eq_f @eq_f >(mconfig_expand ??? outc2)
282 @single_final // @(loop_Some ?????? Hloop2)
284 |cases (HR3 (ctapes sig ?? outc1)) #k2 * #outc2 * #Hloop2 #HM3
285 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
287 [@(loop_merge ?????????
288 (mk_mconfig ? (states sig n (ifTM sig n M1 (single_finalTM … M2) M3 acc)) n
289 (inr (states sig n M1) ? (inr (states sig n (single_finalTM ?? M2)) ? (start sig n M3))) (ctapes ??? outc1) )
292 (lift_confL sig n (states ?? M1) (FinSum (states ?? (single_finalTM … M2)) (states ?? M3)))
293 (step sig n M1) (step sig n (ifTM sig n M1 (single_finalTM ?? M2) M3 acc))
294 (λc.halt sig n M1 (cstate … c))
295 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
298 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
299 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
300 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
302 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1);
303 whd in match (lift_confL ?????);
304 >(trans_if_M1true_notacc … Hacc)
305 [@mconfig_eq // (* whd in ⊢ (??%?); *)
306 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
307 |@(loop_Some ?????? Hloop1)]
308 |cases outc1 #s1 #t1 %
310 (λc.(lift_confR … (lift_confR sig n (states ?? (single_finalTM ?? M2)) (states ?? M3) c)))
313 | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
315 |>(mconfig_expand ??? outc2) whd in match (lift_confR ?????);
318 |#_ @(ex_intro … (ctapes ??? outc1)) %
319 [@HMfalse @(\Pf Hacc) | >(mconfig_expand ??? outc2) @HM3 ]
324 lemma acc_sem_if_app: ∀sig,n,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,R5,acc.
325 M1 ⊨ [acc: Rtrue, Rfalse] → M2 ⊨ R2 → M3 ⊨ R3 →
326 (∀t1,t2,t3. Rtrue t1 t3 → R2 t3 t2 → R4 t1 t2) →
327 (∀t1,t2,t3. Rfalse t1 t3 → R3 t3 t2 → R5 t1 t2) →
328 ifTM sig n M1 (single_finalTM … M2) M3 acc ⊨
329 [inr … (inl … (inr … start_nop)): R4, R5].
330 #sig #n #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #R5 #acc
331 #HRacc #HRtrue #HRfalse #Hsub1 #Hsub2
332 #t cases (acc_sem_if … HRacc HRtrue HRfalse t)
333 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
335 |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
336 |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]
339 lemma sem_single_final_guarded: ∀sig,n.∀M: mTM sig n.∀Pre,R.
340 GRealize sig n M Pre R → GRealize sig n (single_finalTM sig n M) Pre R.
341 #sig #n #M #Pre #R #HR #intape #HPre
342 cases (sem_seq_guarded ???????? HR (Realize_to_GRealize ??? (λt.True) ? (sem_nop …)) ?? HPre) //
343 #k * #outc * #Hloop * #ta * #Hta whd in ⊢ (%→?); #Houtc
344 @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop | >Houtc // ]
347 lemma acc_sem_if_guarded: ∀sig,n.∀M1,M2,M3: mTM sig n.∀P,P2,Rtrue,Rfalse,R2,R3,acc.
348 M1 ⊨ [acc: Rtrue, Rfalse] →
349 (GRealize ?? M2 P2 R2) → (∀t,t0.P t → Rtrue t t0 → P2 t0) →
351 accGRealize ?? (ifTM sig n M1 (single_finalTM … M2) M3 acc)
352 (inr … (inl … (inr … start_nop))) P (Rtrue ∘ R2) (Rfalse ∘ R3).
353 #sig #n #M1 #M2 #M3 #P #P2 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HP2 #HR3 #t #HPt
354 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse
355 cases (true_or_false (cstate ??? outc1 == acc)) #Hacc
356 [lapply (sem_single_final_guarded … HR2) -HR2 #HR2
357 cases (HR2 (ctapes sig ?? outc1) ?)
358 [|@HP2 [||@HMtrue @(\P Hacc)] // ]
359 #k2 * #outc2 * #Hloop2 #HM2
360 @(ex_intro … (k1+k2))
361 @(ex_intro … (lift_confR … (lift_confL … outc2))) %
363 [@(loop_merge ?????????
364 (mk_mconfig ? (states sig n (ifTM sig n M1 (single_finalTM … M2) M3 acc)) n
365 (inr (states sig n M1) ? (inl ? (states sig n M3) (start sig n (single_finalTM sig n M2)))) (ctapes ??? outc1) )
368 (lift_confL sig n (states ?? M1) (FinSum (states ?? (single_finalTM … M2)) (states ?? M3)))
369 (step sig n M1) (step sig n (ifTM sig n M1 (single_finalTM ?? M2) M3 acc))
370 (λc.halt sig n M1 (cstate … c))
371 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
374 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
375 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
376 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
378 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1);
379 whd in match (lift_confL ?????);
380 >(trans_if_M1true_acc … Hacc)
381 [@mconfig_eq // (* whd in ⊢ (??%?); *)
382 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
383 | @(loop_Some ?????? Hloop1)]
384 |cases outc1 #s1 #t1 %
386 (λc.(lift_confR … (lift_confL sig n (states ?? (single_finalTM ?? M2)) (states ?? M3) c)))
389 | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
391 |#_ @(ex_intro … (ctapes ??? outc1)) %
392 [@HMtrue @(\P Hacc) | >(mconfig_expand ??? outc2) @HM2 ]
394 |>(mconfig_expand ??? outc2) whd in match (lift_confR ?????);
395 * #H @False_ind @H @eq_f @eq_f >(mconfig_expand ??? outc2)
396 @single_final // @(loop_Some ?????? Hloop2)
398 |cases (HR3 (ctapes sig ?? outc1)) #k2 * #outc2 * #Hloop2 #HM3
399 @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
401 [@(loop_merge ?????????
402 (mk_mconfig ? (states sig n (ifTM sig n M1 (single_finalTM … M2) M3 acc)) n
403 (inr (states sig n M1) ? (inr (states sig n (single_finalTM ?? M2)) ? (start sig n M3))) (ctapes ??? outc1) )
406 (lift_confL sig n (states ?? M1) (FinSum (states ?? (single_finalTM … M2)) (states ?? M3)))
407 (step sig n M1) (step sig n (ifTM sig n M1 (single_finalTM ?? M2) M3 acc))
408 (λc.halt sig n M1 (cstate … c))
409 (λc.halt_liftL ?? (halt sig n M1) (cstate … c))
412 [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
413 | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
414 |#c0 #Hhalt >(step_if_liftM1 … Hhalt) //
416 |whd in ⊢ (??%?); >(mconfig_expand ??? outc1);
417 whd in match (lift_confL ?????);
418 >(trans_if_M1true_notacc … Hacc)
419 [@mconfig_eq // (* whd in ⊢ (??%?); *)
420 <(tape_move_null_action sig n (ctapes sig (states sig n M1) n outc1)) in ⊢ (???%); %
421 | @(loop_Some ?????? Hloop1)]
422 |cases outc1 #s1 #t1 %
424 (λc.(lift_confR … (lift_confR sig n (states ?? (single_finalTM ?? M2)) (states ?? M3) c)))
427 | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
429 |>(mconfig_expand ??? outc2) whd in match (lift_confR ?????);
432 |#_ @(ex_intro … (ctapes ??? outc1)) %
433 [@HMfalse @(\Pf Hacc) | >(mconfig_expand ??? outc2) @HM3 ]
438 lemma acc_sem_if_app_guarded: ∀sig,n,M1,M2,M3,P,P2,Rtrue,Rfalse,R2,R3,R4,R5,acc.
439 M1 ⊨ [acc: Rtrue, Rfalse] →
440 (GRealize ? n M2 P2 R2) → (∀t,t0.P t → Rtrue t t0 → P2 t0) →
442 (∀t1,t2,t3. Rtrue t1 t3 → R2 t3 t2 → R4 t1 t2) →
443 (∀t1,t2,t3. Rfalse t1 t3 → R3 t3 t2 → R5 t1 t2) →
444 accGRealize ? n (ifTM sig n M1 (single_finalTM … M2) M3 acc)
445 (inr … (inl … (inr … start_nop))) P R4 R5 .
446 #sig #n #M1 #M2 #M3 #P #P2 #Rtrue #Rfalse #R2 #R3 #R4 #R5 #acc
447 #HRacc #HRtrue #Hinv #HRfalse #Hsub1 #Hsub2
448 #t #HPt cases (acc_sem_if_guarded … HRacc HRtrue Hinv HRfalse t HPt)
449 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
451 |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
452 |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]