]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/if_machine.ma
made executable again
[helm.git] / matita / matita / lib / turing / if_machine.ma
1 (*
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.           
5     ||I||                                                            
6     ||T||  
7     ||A||  
8     \   /  This file is distributed under the terms of the       
9      \ /   GNU General Public License Version 2   
10       V_____________________________________________________________*)
11
12 include "turing/mono.ma".
13
14 (**************************** single final machine ****************************)
15
16 definition single_finalTM ≝ 
17   λsig.λM:TM sig.seq ? M (nop ?).
18
19 lemma sem_single_final: ∀sig.∀M: TM sig.∀R.
20   M ⊨ R → single_finalTM sig M ⊨ R.
21 #sig #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 // ]
25 qed.
26
27 lemma single_final: ∀sig.∀M: TM sig.∀q1,q2.
28   halt ? (single_finalTM sig M) q1 = true 
29     →  halt ? (single_finalTM sig M) q2 = true → q1=q2.
30 #sig #M * 
31   [#q1M #q2 whd in match (halt ???); #H destruct
32   |#q1nop *
33     [#q2M #_ whd in match (halt ???); #H destruct
34     |#q2nop #_ #_ @eq_f normalize @nop_single_state
35     ]
36   ]
37 qed.
38   
39 (******************************** if machine **********************************)
40
41 definition if_trans ≝ λsig. λM1,M2,M3 : TM sig. λq:states sig M1.
42 λp. let 〈s,a〉 ≝ p in
43   match s with 
44   [ inl s1 ⇒ 
45       if halt sig M1 s1 then
46         if s1==q then 〈inr … (inl … (start sig M2)), None ?,N〉
47         else 〈inr … (inr … (start sig M3)), None ?,N〉
48       else let 〈news1,newa,m〉 ≝ trans sig M1 〈s1,a〉 in
49        〈inl … news1,newa,m〉
50   | inr s' ⇒ 
51       match s' with
52       [ inl s2 ⇒ let 〈news2,newa,m〉 ≝ trans sig M2 〈s2,a〉 in
53          〈inr … (inl … news2),newa,m〉
54       | inr s3 ⇒ let 〈news3,newa,m〉 ≝ trans sig M3 〈s3,a〉 in
55          〈inr … (inr … news3),newa,m〉
56       ]
57   ]. 
58  
59 definition ifTM ≝ λsig. λcondM,thenM,elseM : TM sig.
60   λqacc: states sig condM.
61   mk_TM sig 
62     (FinSum (states sig condM) (FinSum (states sig thenM) (states sig elseM)))
63     (if_trans sig condM thenM elseM qacc)
64     (inl … (start sig condM))
65     (λs.match s with
66       [ inl _ ⇒ false 
67       | inr s' ⇒ match s' with 
68         [ inl s2 ⇒ halt sig thenM s2
69         | inr s3 ⇒ halt sig elseM s3 ]]).
70
71 (****************************** lifting lemmas ********************************)
72 lemma trans_if_liftM1 : ∀sig,M1,M2,M3,acc,s,a,news,newa,move.
73   halt ? M1 s = false → 
74   trans sig M1 〈s,a〉 = 〈news,newa,move〉 → 
75   trans sig (ifTM sig M1 M2 M3 acc) 〈inl … s,a〉 = 〈inl … news,newa,move〉.
76 #sig * #Q1 #T1 #init1 #halt1 #M2 #M3 #acc #s #a #news #newa #move
77 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
78 qed.
79
80 lemma trans_if_liftM2 : ∀sig,M1,M2,M3,acc,s,a,news,newa,move.
81   halt ? M2 s = false → 
82   trans sig M2 〈s,a〉 = 〈news,newa,move〉 → 
83   trans sig (ifTM sig M1 M2 M3 acc) 〈inr … (inl … s),a〉 = 〈inr… (inl … news),newa,move〉.
84 #sig #M1 * #Q2 #T2 #init2 #halt2 #M3 #acc #s #a #news #newa #move
85 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
86 qed.
87
88 lemma trans_if_liftM3 : ∀sig,M1,M2,M3,acc,s,a,news,newa,move.
89   halt ? M3 s = false → 
90   trans sig M3 〈s,a〉 = 〈news,newa,move〉 → 
91   trans sig (ifTM sig M1 M2 M3 acc) 〈inr … (inr … s),a〉 = 〈inr… (inr … news),newa,move〉.
92 #sig #M1 * #Q2 #T2 #init2 #halt2 #M3 #acc #s #a #news #newa #move
93 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
94 qed.
95
96 lemma step_if_liftM1 : ∀sig,M1,M2,M3,acc,c0.
97  halt ? M1 (cstate ?? c0) = false → 
98  step sig (ifTM sig M1 M2 M3 acc) (lift_confL sig (states ? M1) ? c0) =
99  lift_confL sig (states ? M1) ? (step sig M1 c0).
100 #sig #M1 #M2 #M3 #acc * #s #t
101   lapply (refl ? (trans ?? 〈s,current sig t〉))
102   cases (trans ?? 〈s,current sig t〉) in ⊢ (???% → %);
103   * #s0 #a0 #m0 cases t
104   [ #Heq #Hhalt
105   | 2,3: #s1 #l1 #Heq #Hhalt 
106   |#ls #s1 #rs #Heq #Hhalt ]
107   whd in ⊢ (???(????%)); >Heq whd in ⊢ (???%);
108   whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_if_liftM1 … Hhalt Heq) //
109 qed.
110
111 lemma step_if_liftM2 : ∀sig,M1,M2,M3,acc,c0.
112  halt ? M2 (cstate ?? c0) = false → 
113  step sig (ifTM sig M1 M2 M3 acc) (lift_confR sig ?? (lift_confL sig ?? c0)) =
114  lift_confR sig ?? (lift_confL sig ?? (step sig M2 c0)).
115 #sig #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 #M3 #acc * #s #t
116   lapply (refl ? (trans ?? 〈s,current sig t〉))
117   cases (trans ?? 〈s,current sig t〉) in ⊢ (???% → %);
118   * #s0 #a0 #m0 cases t
119   [ #Heq #Hhalt
120   | 2,3: #s1 #l1 #Heq #Hhalt 
121   |#ls #s1 #rs #Heq #Hhalt ] 
122   whd in match (step ? M2 ?); >Heq whd in ⊢ (???%);
123   whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_if_liftM2 … Hhalt Heq) //
124 qed.
125
126 lemma step_if_liftM3 : ∀sig,M1,M2,M3,acc,c0.
127  halt ? M3 (cstate ?? c0) = false → 
128  step sig (ifTM sig M1 M2 M3 acc) (lift_confR sig ?? (lift_confR sig ?? c0)) =
129  lift_confR sig ?? (lift_confR sig ?? (step sig M3 c0)).
130 #sig #M1 (* * #Q1 #T1 #init1 #halt1 *) #M2 #M3 #acc * #s #t
131   lapply (refl ? (trans ?? 〈s,current sig t〉))
132   cases (trans ?? 〈s,current sig t〉) in ⊢ (???% → %);
133   * #s0 #a0 #m0 cases t
134   [ #Heq #Hhalt
135   | 2,3: #s1 #l1 #Heq #Hhalt 
136   |#ls #s1 #rs #Heq #Hhalt ] 
137   whd in match (step ? M3 ?); >Heq whd in ⊢ (???%);
138   whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_if_liftM3 … Hhalt Heq) //
139 qed.
140
141 lemma trans_if_M1true_acc : ∀sig,M1,M2,M3,acc,s,a.
142   halt ? M1 s = true → s==acc = true →
143   trans sig (ifTM sig M1 M2 M3 acc) 〈inl … s,a〉 = 〈inr … (inl … (start ? M2)),None ?,N〉.
144 #sig #M1 #M2 #M3 #acc #s #a #Hhalt #Hacc whd in ⊢ (??%?); >Hhalt >Hacc %
145 qed.
146
147 lemma trans_if_M1true_notacc : ∀sig,M1,M2,M3,acc,s,a.
148   halt ? M1 s = true → s==acc = false →
149   trans sig (ifTM sig M1 M2 M3 acc) 〈inl … s,a〉 = 〈inr … (inr … (start ? M3)),None ?,N〉.
150 #sig #M1 #M2 #M3 #acc #s #a #Hhalt #Hacc whd in ⊢ (??%?); >Hhalt >Hacc %
151 qed.
152
153 (******************************** semantics ***********************************)
154 lemma sem_if: ∀sig.∀M1,M2,M3:TM sig.∀Rtrue,Rfalse,R2,R3,acc.
155   M1 ⊨ [acc: Rtrue,Rfalse] → M2 ⊨ R2 → M3 ⊨ R3 → 
156     ifTM sig M1 M2 M3 acc ⊨ (Rtrue ∘ R2) ∪ (Rfalse ∘ R3).
157 #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HR3 #t 
158 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse 
159 cases (true_or_false (cstate ?? outc1 == acc)) #Hacc
160   [cases (HR2 (ctape sig ? outc1)) #k2 * #outc2 * #Hloop2 #HM2
161    @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confL … outc2))) %
162     [@(loop_merge ?????????
163      (mk_config ? (FinSum (states sig M1) (FinSum (states sig M2) (states sig M3)))
164       (inr (states sig M1) ? (inl (states sig M2) (states sig M3) (start sig M2))) (ctape ?? outc1) )
165      ? 
166      (loop_lift ??? 
167        (lift_confL sig (states ? M1) (FinSum (states ? M2) (states ? M3)))
168        (step sig M1) (step sig (ifTM sig M1 M2 M3 acc)) 
169        (λc.halt sig M1 (cstate … c)) 
170        (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
171        … Hloop1))
172       [* *
173         [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
174         | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
175       |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
176       |#x <p_halt_liftL %
177       |whd in ⊢ (??%?); >(config_expand ?? outc1);
178        whd in match (lift_confL ????);
179        >(trans_if_M1true_acc … Hacc) 
180         [% | @(loop_Some ?????? Hloop1)]
181       |cases outc1 #s1 #t1 %
182       |@(loop_lift ??? 
183          (λc.(lift_confR … (lift_confL sig (states ? M2) (states ? M3) c)))
184          … Hloop2) 
185         [ * #s2 #t2 %
186         | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
187       ]
188     |%1 @(ex_intro … (ctape ?? outc1)) % 
189       [@HMtrue @(\P Hacc) | >(config_expand ?? outc2) @HM2 ]
190     ]
191   |cases (HR3 (ctape sig ? outc1)) #k2 * #outc2 * #Hloop2 #HM3
192    @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
193     [@(loop_merge ?????????
194      (mk_config ? (FinSum (states sig M1) (FinSum (states sig M2) (states sig M3)))
195       (inr (states sig M1) ? (inr (states sig M2) (states sig M3) (start sig M3))) (ctape ?? outc1) )
196      ? 
197      (loop_lift ??? 
198        (lift_confL sig (states ? M1) (FinSum (states ? M2) (states ? M3)))
199        (step sig M1) (step sig (ifTM sig M1 M2 M3 acc)) 
200        (λc.halt sig M1 (cstate … c)) 
201        (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
202        … Hloop1))
203       [* *
204         [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
205         | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
206       |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
207       |#x <p_halt_liftL %
208       |whd in ⊢ (??%?); >(config_expand ?? outc1);
209        whd in match (lift_confL ????);
210        >(trans_if_M1true_notacc … Hacc) 
211         [% | @(loop_Some ?????? Hloop1)]
212       |cases outc1 #s1 #t1 %
213       |@(loop_lift ??? 
214          (λc.(lift_confR … (lift_confR sig (states ? M2) (states ? M3) c)))
215          … Hloop2) 
216         [ * #s2 #t2 %
217         | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
218       ]
219     |%2 @(ex_intro … (ctape ?? outc1)) % 
220       [@HMfalse @(\Pf Hacc) | >(config_expand ?? outc2) @HM3 ]
221     ]
222   ]
223 qed.
224
225 lemma sem_if_app: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,acc.
226   accRealize sig M1 acc Rtrue Rfalse  → M2 ⊨ R2  → M3 ⊨ R3 →  
227     (∀t1,t2,t3. (Rtrue t1 t3 ∧ R2 t3 t2) ∨ (Rfalse t1 t3 ∧ R3 t3 t2) → R4 t1 t2) → 
228     ifTM sig M1 M2 M3 acc ⊨ R4.
229 #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #acc
230 #HRacc #HRtrue #HRfalse #Hsub
231 #t cases (sem_if … HRacc HRtrue HRfalse t)
232 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
233 % [@Hloop] cases Houtc
234   [* #t3 * #Hleft #Hright @(Hsub … t3) %1 /2/
235   |* #t3 * #Hleft #Hright @(Hsub … t3) %2 /2/ ]
236 qed.
237
238 (* weak 
239 lemma sem_if_app: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,acc.
240   accRealize sig M1 acc Rtrue Rfalse  → M2 ⊨ R2  → M3 ⊨ R3 →  
241     (∀t1,t2,t3. (Rtrue t1 t3 → R2 t3 t2) ∨ (Rfalse t1 t3 → R3 t3 t2) → R4 t1 t2) → 
242     ifTM sig M1 M2 M3 acc ⊨ R4.
243 #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #acc
244 #HRacc #HRtrue #HRfalse #Hsub
245 #t cases (sem_if … HRacc HRtrue HRfalse t)
246 #k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
247 % [@Hloop] cases Houtc
248   [* #t3 * #Hleft #Hright @(Hsub … t3) %1 /2/
249   |* #t3 * #Hleft #Hright @(Hsub … t3) %2 /2/ ]
250 qed.
251 *)
252
253 (* we can probably use acc_sem_if to prove sem_if *)
254 (* for sure we can use acc_sem_if_guarded to prove acc_sem_if *)
255 lemma acc_sem_if: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,acc.
256   M1 ⊨ [acc: Rtrue, Rfalse]  → M2 ⊨ R2 → M3 ⊨ R3 → 
257   ifTM sig M1 (single_finalTM … M2) M3 acc ⊨
258      [inr … (inl … (inr … start_nop)): Rtrue ∘ R2, Rfalse ∘ R3].
259 #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HR3 #t 
260 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse 
261 cases (true_or_false (cstate ?? outc1 == acc)) #Hacc
262   [lapply (sem_single_final … HR2) -HR2 #HR2
263    cases (HR2 (ctape sig ? outc1)) #k2 * #outc2 * #Hloop2 #HM2
264    @(ex_intro … (k1+k2)) 
265    @(ex_intro … (lift_confR … (lift_confL … outc2))) %
266     [%
267       [@(loop_merge ?????????
268          (mk_config ? (states sig (ifTM sig M1 (single_finalTM … M2) M3 acc))
269           (inr (states sig M1) ? (inl ? (states sig M3) (start sig (single_finalTM sig M2)))) (ctape ?? outc1) )
270          ? 
271          (loop_lift ??? 
272           (lift_confL sig (states ? M1) (FinSum (states ? (single_finalTM … M2)) (states ? M3)))
273           (step sig M1) (step sig (ifTM sig M1 (single_finalTM ? M2) M3 acc)) 
274           (λc.halt sig M1 (cstate … c)) 
275           (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
276           … Hloop1))
277         [* *
278           [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
279           | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
280         |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
281         |#x <p_halt_liftL %
282         |whd in ⊢ (??%?); >(config_expand ?? outc1);
283          whd in match (lift_confL ????);
284          >(trans_if_M1true_acc … Hacc) 
285           [% | @(loop_Some ?????? Hloop1)]
286         |cases outc1 #s1 #t1 %
287         |@(loop_lift ??? 
288            (λc.(lift_confR … (lift_confL sig (states ? (single_finalTM ? M2)) (states ? M3) c)))
289            … Hloop2) 
290           [ * #s2 #t2 %
291           | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
292         ]
293       |#_ @(ex_intro … (ctape ?? outc1)) % 
294         [@HMtrue @(\P Hacc) | >(config_expand ?? outc2) @HM2 ]
295       ]
296     |>(config_expand ?? outc2) whd in match (lift_confR ????);
297      * #H @False_ind @H @eq_f @eq_f >(config_expand ?? outc2)
298      @single_final // @(loop_Some ?????? Hloop2)
299     ]
300   |cases (HR3 (ctape sig ? outc1)) #k2 * #outc2 * #Hloop2 #HM3
301    @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
302     [%
303       [@(loop_merge ?????????
304          (mk_config ? (states sig (ifTM sig M1 (single_finalTM … M2) M3 acc))
305           (inr (states sig M1) ? (inr (states sig (single_finalTM ? M2)) ? (start sig M3))) (ctape ?? outc1) )
306          ? 
307          (loop_lift ??? 
308           (lift_confL sig (states ? M1) (FinSum (states ? (single_finalTM … M2)) (states ? M3)))
309           (step sig M1) (step sig (ifTM sig M1 (single_finalTM ? M2) M3 acc)) 
310           (λc.halt sig M1 (cstate … c)) 
311           (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
312           … Hloop1))
313         [* *
314           [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
315           | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
316         |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
317         |#x <p_halt_liftL %
318         |whd in ⊢ (??%?); >(config_expand ?? outc1);
319          whd in match (lift_confL ????);
320          >(trans_if_M1true_notacc … Hacc) 
321           [% | @(loop_Some ?????? Hloop1)]
322         |cases outc1 #s1 #t1 %
323         |@(loop_lift ??? 
324            (λc.(lift_confR … (lift_confR sig (states ? (single_finalTM ? M2)) (states ? M3) c)))
325            … Hloop2) 
326           [ * #s2 #t2 %
327           | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
328         ]
329       |>(config_expand ?? outc2) whd in match (lift_confR ????);
330        #H destruct (H) 
331       ]
332     |#_ @(ex_intro … (ctape ?? outc1)) % 
333       [@HMfalse @(\Pf Hacc) | >(config_expand ?? outc2) @HM3 ]
334     ]
335   ]
336 qed.
337     
338 lemma acc_sem_if_app: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,R5,acc.
339   M1 ⊨ [acc: Rtrue, Rfalse] → M2 ⊨ R2 → M3 ⊨ R3 → 
340     (∀t1,t2,t3. Rtrue t1 t3 → R2 t3 t2 → R4 t1 t2) → 
341     (∀t1,t2,t3. Rfalse t1 t3 → R3 t3 t2 → R5 t1 t2) → 
342      ifTM sig M1 (single_finalTM … M2) M3 acc ⊨ 
343        [inr … (inl … (inr … start_nop)): R4, R5].    
344 #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #R5 #acc
345 #HRacc #HRtrue #HRfalse #Hsub1 #Hsub2 
346 #t cases (acc_sem_if … HRacc HRtrue HRfalse t)
347 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
348 % [% [@Hloop
349      |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
350   |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]
351 qed.
352
353 lemma sem_single_final_guarded: ∀sig.∀M: TM sig.∀Pre,R.
354   GRealize sig M Pre R → GRealize sig (single_finalTM sig M) Pre R.
355 #sig #M #Pre #R #HR #intape #HPre 
356 cases (sem_seq_guarded ??????? HR (Realize_to_GRealize ?? (λt.True) ? (sem_nop …)) ?? HPre) //
357 #k * #outc * #Hloop * #ta * #Hta whd in ⊢ (%→?); #Houtc
358 @(ex_intro ?? k) @(ex_intro ?? outc) %  [ @Hloop | >Houtc // ]
359 qed.
360     
361 lemma acc_sem_if_guarded: ∀sig,M1,M2,M3,P,P2,Rtrue,Rfalse,R2,R3,acc.
362   M1 ⊨ [acc: Rtrue, Rfalse]  → 
363   (GRealize ? M2 P2 R2) → (∀t,t0.P t → Rtrue t t0 → P2 t0) → 
364   M3 ⊨ R3 → 
365   accGRealize ? (ifTM sig M1 (single_finalTM … M2) M3 acc) 
366     (inr … (inl … (inr … start_nop))) P (Rtrue ∘ R2) (Rfalse ∘ R3).
367 #sig #M1 #M2 #M3 #P #P2 #Rtrue #Rfalse #R2 #R3 #acc #HaccR #HR2 #HP2 #HR3 #t #HPt 
368 cases (HaccR t) #k1 * #outc1 * * #Hloop1 #HMtrue #HMfalse 
369 cases (true_or_false (cstate ?? outc1 == acc)) #Hacc
370   [lapply (sem_single_final_guarded … HR2) -HR2 #HR2
371    cases (HR2 (ctape sig ? outc1) ?)
372    [|@HP2 [||@HMtrue @(\P Hacc)] // ]
373    #k2 * #outc2 * #Hloop2 #HM2
374    @(ex_intro … (k1+k2)) 
375    @(ex_intro … (lift_confR … (lift_confL … outc2))) %
376     [%
377       [@(loop_merge ?????????
378          (mk_config ? (states sig (ifTM sig M1 (single_finalTM … M2) M3 acc))
379           (inr (states sig M1) ? (inl ? (states sig M3) (start sig (single_finalTM sig M2)))) (ctape ?? outc1) )
380          ? 
381          (loop_lift ??? 
382           (lift_confL sig (states ? M1) (FinSum (states ? (single_finalTM … M2)) (states ? M3)))
383           (step sig M1) (step sig (ifTM sig M1 (single_finalTM ? M2) M3 acc)) 
384           (λc.halt sig M1 (cstate … c)) 
385           (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
386           … Hloop1))
387         [* *
388           [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
389           | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
390         |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
391         |#x <p_halt_liftL %
392         |whd in ⊢ (??%?); >(config_expand ?? outc1);
393          whd in match (lift_confL ????);
394          >(trans_if_M1true_acc … Hacc) 
395           [% | @(loop_Some ?????? Hloop1)]
396         |cases outc1 #s1 #t1 %
397         |@(loop_lift ??? 
398            (λc.(lift_confR … (lift_confL sig (states ? (single_finalTM ? M2)) (states ? M3) c)))
399            … Hloop2) 
400           [ * #s2 #t2 %
401           | #c0 #Hhalt >(step_if_liftM2 … Hhalt) // ]
402         ]
403       |#_ @(ex_intro … (ctape ?? outc1)) % 
404         [@HMtrue @(\P Hacc) | >(config_expand ?? outc2) @HM2 ]
405       ]
406     |>(config_expand ?? outc2) whd in match (lift_confR ????);
407      * #H @False_ind @H @eq_f @eq_f >(config_expand ?? outc2)
408      @single_final // @(loop_Some ?????? Hloop2)
409     ]
410   |cases (HR3 (ctape sig ? outc1)) #k2 * #outc2 * #Hloop2 #HM3
411    @(ex_intro … (k1+k2)) @(ex_intro … (lift_confR … (lift_confR … outc2))) %
412     [%
413       [@(loop_merge ?????????
414          (mk_config ? (states sig (ifTM sig M1 (single_finalTM … M2) M3 acc))
415           (inr (states sig M1) ? (inr (states sig (single_finalTM ? M2)) ? (start sig M3))) (ctape ?? outc1) )
416          ? 
417          (loop_lift ??? 
418           (lift_confL sig (states ? M1) (FinSum (states ? (single_finalTM … M2)) (states ? M3)))
419           (step sig M1) (step sig (ifTM sig M1 (single_finalTM ? M2) M3 acc)) 
420           (λc.halt sig M1 (cstate … c)) 
421           (λc.halt_liftL ?? (halt sig M1) (cstate … c)) 
422           … Hloop1))
423         [* *
424           [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
425           | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
426         |#c0 #Hhalt >(step_if_liftM1 … Hhalt) // 
427         |#x <p_halt_liftL %
428         |whd in ⊢ (??%?); >(config_expand ?? outc1);
429          whd in match (lift_confL ????);
430          >(trans_if_M1true_notacc … Hacc) 
431           [% | @(loop_Some ?????? Hloop1)]
432         |cases outc1 #s1 #t1 %
433         |@(loop_lift ??? 
434            (λc.(lift_confR … (lift_confR sig (states ? (single_finalTM ? M2)) (states ? M3) c)))
435            … Hloop2) 
436           [ * #s2 #t2 %
437           | #c0 #Hhalt >(step_if_liftM3 … Hhalt) // ]
438         ]
439       |>(config_expand ?? outc2) whd in match (lift_confR ????);
440        #H destruct (H) 
441       ]
442     |#_ @(ex_intro … (ctape ?? outc1)) % 
443       [@HMfalse @(\Pf Hacc) | >(config_expand ?? outc2) @HM3 ]
444     ]
445   ]
446 qed.
447     
448 lemma acc_sem_if_app_guarded: ∀sig,M1,M2,M3,P,P2,Rtrue,Rfalse,R2,R3,R4,R5,acc.
449   M1 ⊨ [acc: Rtrue, Rfalse] → 
450   (GRealize ? M2 P2 R2) → (∀t,t0.P t → Rtrue t t0 → P2 t0) → 
451   M3 ⊨ R3 → 
452   (∀t1,t2,t3. Rtrue t1 t3 → R2 t3 t2 → R4 t1 t2) → 
453   (∀t1,t2,t3. Rfalse t1 t3 → R3 t3 t2 → R5 t1 t2) → 
454   accGRealize ? (ifTM sig M1 (single_finalTM … M2) M3 acc) 
455     (inr … (inl … (inr … start_nop))) P R4 R5 .
456 #sig #M1 #M2 #M3 #P #P2 #Rtrue #Rfalse #R2 #R3 #R4 #R5 #acc
457 #HRacc #HRtrue #Hinv #HRfalse #Hsub1 #Hsub2 
458 #t #HPt cases (acc_sem_if_guarded … HRacc HRtrue Hinv HRfalse t HPt)
459 #k * #outc * * #Hloop #Houtc1 #Houtc2 @(ex_intro … k) @(ex_intro … outc)
460 % [% [@Hloop
461      |#H cases (Houtc1 H) #t3 * #Hleft #Hright @Hsub1 // ]
462   |#H cases (Houtc2 H) #t3 * #Hleft #Hright @Hsub2 // ]
463 qed.
464
465