]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/match.ma
6e5dad0525594cdb215cae36289afc210e6dd882
[helm.git] / matita / matita / lib / turing / multi_universal / match.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "turing/multi_universal/compare.ma".
16 include "turing/multi_universal/par_test.ma".
17 include "turing/multi_universal/moves_2.ma".
18
19 definition Rtc_multi_true ≝ 
20   λalpha,test,n,i.λt1,t2:Vector ? (S n).
21    (∃c. current alpha (nth i ? t1 (niltape ?)) = Some ? c ∧ test c = true) ∧ t2 = t1.
22    
23 definition Rtc_multi_false ≝ 
24   λalpha,test,n,i.λt1,t2:Vector ? (S n).
25     (∀c. current alpha (nth i ? t1 (niltape ?)) = Some ? c → test c = false) ∧ t2 = t1.
26
27 lemma sem_test_char_multi :
28   ∀alpha,test,n,i.i ≤ n → 
29   inject_TM ? (test_char ? test) n i ⊨ 
30   [ tc_true : Rtc_multi_true alpha test n i, Rtc_multi_false alpha test n i ].
31 #alpha #test #n #i #Hin #int
32 cases (acc_sem_inject … Hin (sem_test_char alpha test) int)
33 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
34 [ @Hloop
35 | #Hqtrue lapply (Htrue Hqtrue) * * * #c *
36   #Hcur #Htestc #Hnth_i #Hnth_j %
37   [ %{c} % //
38   | @(eq_vec … (niltape ?)) #i0 #Hi0
39     cases (decidable_eq_nat i0 i) #Hi0i
40     [ >Hi0i @Hnth_i
41     | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
42 | #Hqfalse lapply (Hfalse Hqfalse) * * #Htestc #Hnth_i #Hnth_j %
43   [ @Htestc
44   | @(eq_vec … (niltape ?)) #i0 #Hi0
45     cases (decidable_eq_nat i0 i) #Hi0i
46     [ >Hi0i @Hnth_i
47     | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
48 qed.
49
50 definition Rm_test_null_true ≝ 
51   λalpha,n,i.λt1,t2:Vector ? (S n).
52    current alpha (nth i ? t1 (niltape ?)) ≠ None ? ∧ t2 = t1.
53    
54 definition Rm_test_null_false ≝ 
55   λalpha,n,i.λt1,t2:Vector ? (S n).
56     current alpha (nth i ? t1 (niltape ?)) = None ? ∧ t2 = t1.
57
58 lemma sem_test_null_multi : ∀alpha,n,i.i ≤ n → 
59   inject_TM ? (test_null ?) n i ⊨ 
60     [ tc_true : Rm_test_null_true alpha n i, Rm_test_null_false alpha n i ].
61 #alpha #n #i #Hin #int
62 cases (acc_sem_inject … Hin (sem_test_null alpha) int)
63 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
64 [ @Hloop
65 | #Hqtrue lapply (Htrue Hqtrue) * * #Hcur #Hnth_i #Hnth_j % //
66   @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) #Hi0i
67   [ >Hi0i @sym_eq @Hnth_i | @sym_eq @Hnth_j @sym_not_eq // ] ] 
68 | #Hqfalse lapply (Hfalse Hqfalse) * * #Hcur #Hnth_i #Hnth_j %
69   [ @Hcur
70   | @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) // 
71     #Hi0i @sym_eq @Hnth_j @sym_not_eq // ] ] 
72 qed.
73
74 definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
75   match (nth src (option sig) v (None ?)) with 
76   [ None ⇒  false 
77   | Some x ⇒  notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
78   
79 definition rewind ≝ λsrc,dst,sig,n.parmove src dst sig n L · parmove_step src dst sig n R.
80
81 definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
82   (∀x,x0,xs,rs.
83     nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
84     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
85     nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
86     outt = change_vec ?? 
87            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
88            (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst).
89            
90 theorem accRealize_to_Realize :
91   ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
92   M ⊨ [ acc: Rtrue, Rfalse ] →  M ⊨ Rtrue ∪ Rfalse.
93 #sig #n #M #Rtrue #Rfalse #acc #HR #t
94 cases (HR t) #k * #outc * * #Hloop
95 #Htrue #Hfalse %{k} %{outc} % // 
96 cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
97 [ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
98 qed. 
99            
100 lemma sem_rewind : ∀src,dst,sig,n.
101   src ≠ dst → src < S n → dst < S n → 
102   rewind src dst sig n ⊨ R_rewind src dst sig n.
103 #src #dst #sig #n #Hneq #Hsrc #Hdst
104 check acc_sem_seq_app
105 @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst)
106   (accRealize_to_Realize … (sem_parmove_step src dst sig n R Hneq Hsrc Hdst)))
107 #ta #tb * #tc * * #HR1 #_ #HR2
108 #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst
109 >(HR1 ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in HR2;
110 [|>Hmidta_dst //
111 |>length_append >length_append >Hlen % ] *
112 [ whd in ⊢ (%→?); * #x1 * #x2 * *
113   >change_vec_commute in ⊢ (%→?); // >nth_change_vec //
114   cases (reverse sig (xs@[x0])@x::rs)
115   [|#z #zs] normalize in ⊢ (%→?); #H destruct (H)
116 | whd in ⊢ (%→?); * #_ #Htb >Htb -Htb FAIL
117   
118    normalize in ⊢ (%→?);
119     (sem_parmove_step src dst sig n R Hneq Hsrc Hdst))
120     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
121       (sem_seq … 
122         (sem_parmoveL ???? Hneq Hsrc Hdst) 
123         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
124       (sem_nop …)))
125   
126
127 definition match_step ≝ λsrc,dst,sig,n.
128   compare src dst sig n ·
129      (ifTM ?? (partest sig n (match_test src dst sig ?))
130       (single_finalTM ??
131         (rewind src dst sig n · (inject_TM ? (move_r ?) n dst)))
132       (nop …)
133       partest1).
134       
135 definition R_match_step_false ≝ 
136   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
137   ∀ls,x,xs.
138   nth src ? int (niltape ?) = midtape sig ls x xs →
139   ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
140     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
141       xs = rs0@xs0 ∧
142       current sig (nth dst (tape sig) outt (niltape sig)) = None ?) ∨
143     (∃ls0,rs0. 
144      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
145      (* ∀rsj,c. 
146      rs0 = c::rsj → *)
147      outt = change_vec ??
148             (change_vec ?? int (mk_tape sig (reverse ? xs@x::ls) (None ?) [ ]) src)
149             (mk_tape sig (reverse ? xs@x::ls0) (option_hd ? rs0) (tail ? rs0)) dst).
150
151 (*definition R_match_step_true ≝ 
152   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
153   ∀s,rs.current sig (nth src (tape sig) int (niltape sig)) = Some ? s → 
154   current sig (nth dst (tape sig) int (niltape sig)) ≠ None ? ∧
155   (∀s1.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s1 → s ≠ s1 →  
156    outt = change_vec ?? int 
157           (tape_move_mono … (nth dst ? int (niltape ?)) (〈Some ? s1,R〉)) dst) ∧  
158   (∀ls,x,xs,ci,rs,ls0,rs0. 
159     nth src ? int (niltape ?) = midtape sig ls x (xs@ci::rs) →
160     nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) →
161     rs0 ≠ [] ∧
162     ∀cj,rs1.rs0 = cj::rs1 → 
163     ci ≠ cj →
164     (outt = change_vec ?? int 
165         (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst)). 
166 *)
167 definition R_match_step_true ≝ 
168   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
169   ∀s.current sig (nth src (tape sig) int (niltape sig)) = Some ? s → 
170   ∃s1.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s1 ∧
171   (left ? (nth src ? int (niltape ?)) = [ ] → 
172    (s ≠ s1 →  
173     outt = change_vec ?? int 
174           (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst) ∧
175    (∀xs,ci,rs,ls0,rs0. 
176     nth src ? int (niltape ?) = midtape sig [] s (xs@ci::rs) →
177     nth dst ? int (niltape ?) = midtape sig ls0 s (xs@rs0) →
178     rs0 ≠ [] ∧
179     ∀cj,rs1.rs0 = cj::rs1 → 
180     ci ≠ cj →
181     (outt = change_vec ?? int 
182         (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst))). 
183          
184 lemma sem_match_step :
185   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
186   match_step src dst sig n ⊨ 
187     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
188       R_match_step_true src dst sig n, 
189       R_match_step_false src dst sig n ].
190 #src #dst #sig #n #Hneq #Hsrc #Hdst 
191 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
192     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
193       (sem_seq … 
194         (sem_parmoveL ???? Hneq Hsrc Hdst) 
195         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
196       (sem_nop …)))
197 [#ta #tb #tc * #Hcomp1 #Hcomp2 * #td * #Htest
198  * #te * #Hte #Htb #s #Hcurta_src whd
199  cut (∃s1.current sig (nth dst (tape sig) ta (niltape sig))=Some sig s1)
200  [ lapply Hcomp1 -Hcomp1 
201    lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
202    cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
203    [ #Hcurta_dst #Hcomp1 >Hcomp1 in Htest; // *
204      change with (vec_map ?????) in match (current_chars ???); whd in ⊢ (??%?→?);
205      <(nth_vec_map ?? (current ?) src ? ta (niltape ?))      
206      <(nth_vec_map ?? (current ?) dst ? ta (niltape ?))
207      >Hcurta_src >Hcurta_dst whd in ⊢ (??%?→?); #H destruct (H)
208    | #s1 #_ #_ %{s1} % ] ]
209  * #s1 #Hcurta_dst %{s1} % // #Hleftta %
210  [ #Hneqss1 -Hcomp2 cut (tc = ta) 
211    [@Hcomp1 %1 %1 >Hcurta_src >Hcurta_dst @(not_to_not … Hneqss1) #H destruct (H) //] 
212     #H destruct (H) -Hcomp1 cut (td = ta)
213     [ cases Htest -Htest // ] #Htdta destruct (Htdta)
214     cases Hte -Hte #Hte #_
215     cases (current_to_midtape … Hcurta_src) #ls * #rs #Hmidta_src
216     cases (current_to_midtape … Hcurta_dst) #ls0 * #rs0 #Hmidta_dst
217     >Hmidta_src in Hleftta; normalize in ⊢ (%→?); #Hls destruct (Hls)
218     >(Hte s [ ] rs Hmidta_src ls0 s1 [ ] rs0 (refl ??) Hmidta_dst) in Htb;
219     * whd in ⊢ (%→?);
220     mid
221     
222       in Htb;
223     cut (te = ta) 
224     [ cases Htest -Htest #Htest #Htdta <Htdta @Hte %1 >Htdta @Hcurta_src %{s} % //] 
225     -Hte #H destruct (H) %
226     [cases Htb * #_ #Hmove #Hmove1 @(eq_vec … (niltape … ))
227      #i #Hi cases (decidable_eq_nat i dst) #Hidst
228       [ >Hidst >nth_change_vec // cases (current_to_midtape … Hcurta_dst)
229         #ls * #rs #Hta_mid >(Hmove … Hta_mid) >Hta_mid cases rs //
230       | >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @Hmove1 @sym_not_eq // ]
231     | whd in Htest:(??%?); >(nth_vec_map ?? (current sig)) in Hcurta_src; #Hcurta_src
232       >Hcurta_src in Htest; whd in ⊢ (??%?→?);
233       cases (is_endc s) // whd in ⊢ (??%?→?); #H @sym_eq // 
234     ]
235    <(nth_vec_map ?? (current ?) dst ? tc (niltape ?))   
236     >Hcurta_src normalize
237    lapply (refl ? (current ? (nth dst ? ta (niltape ?)))) 
238    cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
239    [| #s1 #Hcurta_dst %
240      [ % #Hfalse destruct (Hfalse)
241      | #s1' #Hs1 destruct (Hs1) #Hneqss1 -Hcomp2 
242        cut (tc = ta) 
243        [@Hcomp1 %1 %1 >Hcurta_src >Hcurta_dst @(not_to_not … Hneqss1) #H destruct (H) //] 
244        #H destruct (H) -Hcomp1 cases Hte -Hte #_ #Hte
245        cut (te = ta) [ cases Htest -Htest #Htest #Htdta <Htdta @Hte %1 %{s} % //] -Hte #H destruct (H) %
246        [cases Htb * #_ #Hmove #Hmove1 @(eq_vec … (niltape … ))
247         #i #Hi cases (decidable_eq_nat i dst) #Hidst
248          [ >Hidst >nth_change_vec // cases (current_to_midtape … Hcurta_dst)
249            #ls * #rs #Hta_mid >(Hmove … Hta_mid) >Hta_mid cases rs //
250          | >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @Hmove1 @sym_not_eq // ]
251        | whd in Htest:(??%?); >(nth_vec_map ?? (current sig)) in Hcurta_src; #Hcurta_src
252          >Hcurta_src in Htest; whd in ⊢ (??%?→?);
253          cases (is_endc s) // whd in ⊢ (??%?→?); #H @sym_eq // 
254        ]
255      
256       ]
257    #Hcurta_dst >Hcomp1 in Htest; [| %2 %2 //]
258    whd in ⊢ (??%?→?); change with (current ? (niltape ?)) in match (None ?);
259     <nth_vec_map >Hcurta_src whd in ⊢ (??%?→?); <nth_vec_map
260     >Hcurta_dst cases (is_endc s) normalize in ⊢ (%→?); #H destruct (H)
261  | #Hstart #Hnotstart %
262    [ #s1 #Hcurta_dst #Hneqss1 -Hcomp2
263      cut (tc = ta) 
264      [@Hcomp1 %2 %1 %1 >Hcurta_src >Hcurta_dst @(not_to_not … Hneqss1) #H destruct (H) //] 
265      #H destruct (H) -Hcomp1 cases Hte #_ -Hte #Hte
266      cut (te = ta) [@Hte %1 %1 %{s} % //] -Hte #H destruct (H) %
267      [cases Htb * #_ #Hmove #Hmove1 @(eq_vec … (niltape … ))
268       #i #Hi cases (decidable_eq_nat i dst) #Hidst
269        [ >Hidst >nth_change_vec // cases (current_to_midtape … Hcurta_dst)
270          #ls * #rs #Hta_mid >(Hmove … Hta_mid) >Hta_mid cases rs //
271        | >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @Hmove1 @sym_not_eq // ]
272      | whd in Htest:(??%?); >(nth_vec_map ?? (current sig)) in Hcurta_src; #Hcurta_src
273        >Hcurta_src in Htest; whd in ⊢ (??%?→?);
274        cases (is_endc s) // whd in ⊢ (??%?→?); #H @sym_eq // 
275      ]
276    |#ls #x #xs #ci #rs #ls0 #rs00 #Htasrc_mid #Htadst_mid #Hnotendc
277     cases (Hcomp2 … Htasrc_mid Htadst_mid Hnotendc)
278     [ * #Hrs00 #Htc >Htc in Htest; whd in ⊢ (??%?→?);
279       <(nth_vec_map ?? (current sig) ??? (niltape ?))
280       >change_vec_commute // >nth_change_vec // whd in ⊢ (??%?→?);
281       cases (is_endc ci)
282       [ whd in ⊢ (??%?→?); #H destruct (H)
283       | <(nth_vec_map ?? (current sig) ??? (niltape ?))
284         >change_vec_commute [| @sym_not_eq // ] >nth_change_vec //
285         >(?:current ? (mk_tape ?? (None ?) ?) = None ?)
286         [ whd in ⊢ (??%?→?); #H destruct (H)
287         | cases (reverse sig xs@x::ls0) normalize // ] ] ]
288     * #cj' * #rs0' * #Hcjrs0 destruct (Hcjrs0) -Hcomp2 #Hcomp2 % [ % 
289     [ cases (true_or_false (is_endc ci)) //
290       #Hendci >(Hcomp2 (or_introl … Hendci)) in Htest;
291       whd in ⊢ (??%?→?); <(nth_vec_map ?? (current sig) ??? (niltape ?))
292         >change_vec_commute // >nth_change_vec // whd in ⊢ (??%?→?);
293         >Hendci normalize //
294     | % #H destruct (H) ] ] #cj #rs1 #H destruct (H) #Hcicj
295     lapply (Hcomp2 (or_intror ?? Hcicj)) -Hcomp2 #Htc %
296     [ cases Hte -Hte #Hte #_ whd in Hte;
297       >Htasrc_mid in Hcurta_src; whd in ⊢ (??%?→?); #H destruct (H) 
298       lapply (Hte ls ci (reverse ? xs) rs s ??? ls0 cj (reverse ? xs) s rs1 (refl ??) ?) //
299       [ >Htc >nth_change_vec //
300       | #c0 #Hc0 @(Hnotstart c0) >Htasrc_mid cases (orb_true_l … Hc0) -Hc0 #Hc0
301         [@memb_append_l2 >(\P Hc0) @memb_hd
302         |@memb_append_l1 <(reverse_reverse …xs) @memb_reverse //
303         ]
304       | >Htc >change_vec_commute // >nth_change_vec // ] -Hte
305       >Htc >change_vec_commute // >change_vec_change_vec 
306       >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
307       >Hte in Htb; * * #_ >reverse_reverse #Htbdst1 #Htbdst2 -Hte @(eq_vec … (niltape ?))
308       #i #Hi cases (decidable_eq_nat i dst) #Hidst
309       [ >Hidst >nth_change_vec // >(Htbdst1 ls0 s (xs@cj::rs1))
310         [| >nth_change_vec // ]
311         >Htadst_mid cases xs //
312       | >nth_change_vec_neq [|@sym_not_eq // ]
313         <Htbdst2 [| @sym_not_eq // ] >nth_change_vec_neq [| @sym_not_eq // ]
314         <Htasrc_mid >change_vec_same % ]
315     | >Hcurta_src in Htest; whd in ⊢(??%?→?);
316       >Htc >change_vec_commute //
317       change with (current ? (niltape ?)) in match (None ?);
318       <nth_vec_map >nth_change_vec // whd in ⊢ (??%?→?);
319       cases (is_endc ci) whd in ⊢ (??%?→?); #H destruct (H) % 
320     ]
321    ]
322   ]
323 |#intape #outtape #ta * #Hcomp1 #Hcomp2 * #tb * * #Hc #Htb 
324  whd in ⊢ (%→?); #Hout >Hout >Htb whd
325  #ls #c_src #xs #end #rs #Hmid_src #Hnotend #Hend
326  lapply (current_to_midtape sig (nth dst ? intape (niltape ?)))
327  cases (current … (nth dst ? intape (niltape ?))) in Hcomp1;
328   [#Hcomp1 #_ %1 % % [% | @Hcomp1 %2 %2 % ]
329   |#c_dst cases (true_or_false (c_src == c_dst)) #Hceq
330     [#_ #Hmid_dst cases (Hmid_dst c_dst (refl …)) -Hmid_dst
331      #ls_dst * #rs_dst #Hmid_dst
332      cases (comp_list … (xs@end::rs) rs_dst is_endc) #xs1 * #rsi * #rsj * * * 
333      #Hrs_src #Hrs_dst #Hnotendxs1 #Hneq >Hrs_dst in Hmid_dst; #Hmid_dst
334      cut (∃r1,rs1.rsi = r1::rs1) 
335      [cases rsi in Hrs_src;
336        [ >append_nil #H <H in Hnotendxs1; #Hnotendxs1
337          >(Hnotendxs1 end) in Hend; [ #H1 destruct (H1) ]
338          @memb_append_l2 @memb_hd
339        | #r1 #rs1 #_ %{r1} %{rs1} % ] ]
340      * #r1 * #rs1 #Hrs1 >Hrs1 in Hrs_src;
341      #Hrs_src >Hrs_src in Hmid_src; #Hmid_src <(\P Hceq) in Hmid_dst; #Hmid_dst
342      lapply (Hcomp2 ??????? Hmid_src Hmid_dst ?) 
343      [ #c0 #Hc0 cases (orb_true_l … Hc0) -Hc0 #Hc0
344        [ >(\P Hc0) @Hnotend @memb_hd | @Hnotendxs1 //] ] 
345      *
346      [ * #Hrsj >Hrsj #Hta % %2 >Hta >nth_change_vec //
347        %{ls_dst} %{xs1} cut (∃xs0.xs = xs1@xs0)
348        [lapply Hnotendxs1 -Hnotendxs1 lapply Hrs_src lapply xs elim xs1
349          [ #l #_ #_ %{l} %
350          | #x2 #xs2 #IH * 
351            [ whd in ⊢ (??%%→?); #H destruct (H) #Hnotendxs2
352              >Hnotendxs2 in Hend; [ #H destruct (H) |@memb_hd ]
353            | #x2' #xs2' whd in ⊢ (??%%→?); #H destruct (H)
354              #Hnotendxs2 cases (IH xs2' e0 ?)
355              [ #xs0 #Hxs2 %{xs0} @eq_f //
356              |#c #Hc @Hnotendxs2 @memb_cons // ]
357            ]
358          ] 
359        ] * #xs0 #Hxs0 %{xs0} % [ %
360        [ >Hmid_dst >Hrsj >append_nil %
361        | @Hxs0 ]
362        | cases (reverse ? xs1) // ]
363      | * #cj * #rs2 * #Hrsj #Hta lapply (Hta ?)
364        [ cases (Hneq ?? Hrs1) /2/ * #_ #Hr1 %2 @(Hr1 ?? Hrsj) ] -Hta #Hta
365        %2 >Hta in Hc; whd in ⊢ (??%?→?);
366        change with (current ? (niltape ?)) in match (None ?);
367        <nth_vec_map >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
368        whd in ⊢ (??%?→?); #Hc cut (is_endc r1 = true)
369        [ cases (is_endc r1) in Hc; whd in ⊢ (??%?→?); //
370          change with (current ? (niltape ?)) in match (None ?);
371          <nth_vec_map >nth_change_vec // normalize #H destruct (H) ]
372        #Hendr1 cut (xs = xs1)
373        [ lapply Hnotendxs1 lapply Hnotend lapply Hrs_src lapply xs1
374          -Hnotendxs1 -Hnotend -Hrs_src -xs1 elim xs
375          [ * normalize in ⊢ (%→?); //
376            #x2 #xs2 normalize in ⊢ (%→?); #Heq destruct (Heq) #_ #Hnotendxs1
377            lapply (Hnotendxs1 ? (memb_hd …)) >Hend #H destruct (H)
378          | #x2 #xs2 #IH *
379            [ normalize in ⊢ (%→?); #Heq destruct (Heq) #Hnotendc
380              >Hnotendc in Hendr1; [| @memb_cons @memb_hd ]
381              normalize in ⊢ (%→?); #H destruct (H)
382            | #x3 #xs3 normalize in ⊢ (%→?); #Heq destruct (Heq)
383              #Hnotendc #Hnotendcxs1 @eq_f @IH
384              [ @(cons_injective_r … Heq)
385              | #c0 #Hc0 @Hnotendc cases (orb_true_l … Hc0) -Hc0 #Hc0
386                [ >(\P Hc0) @memb_hd
387                | @memb_cons @memb_cons // ]
388              | #c #Hc @Hnotendcxs1 @memb_cons // ]
389            ]
390          ]
391        | #Hxsxs1 destruct (Hxsxs1) >Hmid_dst %{ls_dst} %{rsj} % //
392          #rsj0 #c >Hrsj #Hrsj0 destruct (Hrsj0) 
393          lapply (append_l2_injective … Hrs_src) // #Hrs' destruct (Hrs') %
394        ]
395      ]
396     |#Hcomp1 #Hsrc cases (Hsrc ? (refl ??)) -Hsrc #ls0 * #rs0 #Hdst 
397      @False_ind lapply (Hcomp1 ?) [%2 %1 %1 >Hmid_src normalize
398      @(not_to_not ??? (\Pf Hceq)) #H destruct //] #Hintape >Hintape in Hc;
399      whd in ⊢(??%?→?); >Hmid_src  
400      change with (current ? (niltape ?)) in match (None ?);
401      <nth_vec_map >Hmid_src whd in ⊢ (??%?→?);
402      >(Hnotend c_src) [|@memb_hd]
403      change with (current ? (niltape ?)) in match (None ?);
404      <nth_vec_map >Hmid_src whd in ⊢ (??%?→?); >Hdst normalize #H destruct (H)
405    ]
406   ]
407 ]
408 qed.
409
410 definition match_m ≝ λsrc,dst,sig,n,is_startc,is_endc.
411   whileTM … (match_step src dst sig n is_startc is_endc) 
412     (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
413
414 definition R_match_m ≝ 
415   λsrc,dst,sig,n,is_startc,is_endc.λint,outt: Vector (tape sig) (S n).
416   ∀ls,x,xs,end,rs.
417   nth src ? int (niltape ?) = midtape sig ls x (xs@end::rs) →
418   (∀c0. memb ? c0 (x::xs) = true → is_endc c0 = false) → is_endc end = true →
419   (∀c0. memb ? c0 (xs@end::rs) = true → is_startc c0 = false) → 
420   (current sig (nth dst (tape sig) int (niltape sig)) = None ? → outt = int) ∧
421   (is_startc x = true →
422    (∀ls0,x0,rs0.
423     nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
424     (∃l,l1.x0::rs0 = l@x::xs@l1 ∧
425      ∀cj,l2.l1=cj::l2 →
426      outt = change_vec ?? 
427             (change_vec ?? int (midtape sig (reverse ? xs@x::ls) end rs) src)
428             (midtape sig ((reverse ? (l@x::xs))@ls0) cj l2) dst) ∨
429     ∀l,l1.x0::rs0 ≠ l@x::xs@l1)).
430
431 lemma not_sub_list_merge : 
432   ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
433 #T #a #b #H1 #H2 #l elim l normalize //
434 qed.
435
436 lemma not_sub_list_merge_2 : 
437   ∀T:DeqSet.∀a,b:list T.∀t. (∀l1.t::a ≠ b@l1) → (∀l,l1.a ≠ l@b@l1) → ∀l,l1.t::a ≠ l@b@l1.
438 #T #a #b #t #H1 #H2 #l elim l //
439 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
440 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
441 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
442 qed.
443
444
445 lemma wsem_match_m : ∀src,dst,sig,n,is_startc,is_endc.
446 src ≠ dst → src < S n → dst < S n → 
447   match_m src dst sig n is_startc is_endc ⊫ R_match_m src dst sig n is_startc is_endc.
448 #src #dst #sig #n #is_startc #is_endc #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
449 lapply (sem_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst) … Hloop) //
450 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
451 [ #Hfalse #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart
452   cases (Hfalse … Hmid_src Hnotend Hend) -Hfalse 
453   [(* current dest = None *) *
454     [ * #Hcur_dst #Houtc %
455       [#_ >Houtc //
456       |#Hstart #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst; 
457        normalize in ⊢ (%→?); #H destruct (H)
458       ]
459     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
460       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
461       | #Hstart #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
462         >Hrs0 cases xs0
463         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
464           #cj #ls2 #H destruct (H)
465         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
466           >length_append whd in ⊢ (??%(??%)→?); >length_append
467           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
468           #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
469           >associative_plus >associative_plus 
470           #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
471           #e2 destruct (e2)
472         ]
473       ]
474     ]
475   |* #ls0 * #rs0 * #Hmid_dst #HFalse %
476     [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
477     | #Hstart #ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
478      %1 %{[ ]} %{rs0} % [%] #cj #l2 #Hnotnil 
479      >reverse_cons >associative_append @(HFalse ?? Hnotnil)
480     ]
481   ]
482 |-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
483  #ls #x #xs #end #rs #Hmid_src #Hnotend #Hend #Hnotstart 
484  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
485  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
486   [#Hmid_dst % 
487     [#_ whd in Htrue; >Hmid_src in Htrue; #Htrue
488      cases (Htrue x (refl … )) -Htrue * #Htaneq #_
489      @False_ind >Hmid_dst in Htaneq; /2/
490     |#Hstart #ls0 #x0 #rs0 #Hmid_dst2 >Hmid_dst2 in Hmid_dst; normalize in ⊢ (%→?); 
491      #H destruct (H)
492     ]
493   | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
494     #Hstart #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcurta_dst; normalize in ⊢ (%→?);
495     #H destruct (H) whd in Htrue; >Hmid_src in Htrue; #Htrue
496     cases (Htrue x (refl …)) -Htrue #_ #Htrue cases (Htrue Hstart Hnotstart) -Htrue
497     cases (true_or_false (x==c)) #eqx
498     [ lapply (\P eqx) -eqx #eqx destruct (eqx)
499       #_ #Htrue cases (comp_list ? (xs@end::rs) rs0 is_endc)
500       #x1 * #tl1 * #tl2 * * * #Hxs #Hrs0 #Hnotendx1
501       cases tl1 in Hxs; 
502       [>append_nil #Hx1 <Hx1 in Hnotendx1; #Hnotendx1
503        lapply (Hnotendx1 end ?) [ @memb_append_l2 @memb_hd ]
504        >Hend #H destruct (H) ]
505       #ci -tl1 #tl1 #Hxs #H cases (H … (refl … )) -H
506       [ #Hendci % >Hrs0 in Hmid_dst; cut (ci = end ∧ x1 = xs) 
507         [ lapply Hxs lapply Hnotendx1 lapply x1 elim xs in Hnotend;
508           [ #_ *
509             [ #_ normalize #H destruct (H) /2/ 
510             | #x2 #xs2 #Hnotendx2 normalize #H destruct (H)
511               >(Hnotendx2 ? (memb_hd …)) in Hend; #H destruct (H) ]
512           | #x2 #xs2 #IH #Hnotendx2 *
513             [ #_ normalize #H destruct (H) >(Hnotendx2 ci ?) in Hendci;
514               [ #H destruct (H)
515               | @memb_cons @memb_hd ]
516             | #x3 #xs3 #Hnotendx3 normalize #H destruct (H)
517               cases (IH … e0)
518               [ #H1 #H2 /2/
519               | #c0 #Hc0 @Hnotendx2 cases (orb_true_l … Hc0) -Hc0 #Hc0
520                 [ >(\P Hc0) @memb_hd
521                 | @memb_cons @memb_cons @Hc0 ]
522               | #c0 #Hc0 @Hnotendx3 @memb_cons @Hc0 ]
523             ]
524           ]
525         | * #Hcieq #Hx1eq >Hx1eq #Hmid_dst 
526           cases (Htrue ??????? (refl ??) Hmid_dst Hnotend)
527           <Hcieq >Hendci * #H destruct (H) ]
528       |cases tl2 in Hrs0;
529         [ >append_nil #Hrs0 destruct (Hrs0) * #Hcifalse#_ %2
530           cut (∃l.xs = x1@ci::l) 
531           [lapply Hxs lapply Hnotendx1 lapply Hnotend lapply xs 
532            -Hxs -xs -Hnotendx1 elim x1
533             [ *
534               [ #_ #_ normalize #H1 destruct (H1) >Hend in Hcifalse;
535                 #H1 destruct (H1)
536               | #x2 #xs2 #_ #_ normalize #H >(cons_injective_l ????? H) %{xs2} % ]
537             | #x2 #xs2 #IHin *
538               [ #_ #Hnotendxs2 normalize #H destruct (H) 
539                 >(Hnotendxs2 ? (memb_hd …)) in Hend; #H destruct (H)
540               | #x3 #xs3 #Hnotendxs3 #Hnotendxs2 normalize #H destruct (H)
541                 cases (IHin ??? e0)
542                 [ #xs4 #Hxs4 >Hxs4 %{xs4} %
543                 | #c0 #Hc0 cases (orb_true_l … Hc0) -Hc0 #Hc0
544                   [ >(\P Hc0) @Hnotendxs3 @memb_hd
545                   | @Hnotendxs3 @memb_cons @memb_cons @Hc0 ]
546                 | #c0 #Hc0 @Hnotendxs2 @memb_cons @Hc0 ]
547               ]
548             ]
549           ] * #l #Hxs' >Hxs'
550           #l0 #l1 % #H lapply (eq_f ?? (length ?) ?? H) -H
551           >length_append normalize >length_append >length_append
552           normalize >commutative_plus normalize #H destruct (H) -H
553           >associative_plus in e0; >associative_plus
554           >(plus_n_O (|x1|)) in ⊢(??%?→?); #H lapply (injective_plus_r … H)
555           -H normalize #H destruct (H)
556         | #cj #tl2' #Hrs0 * #Hcifalse #Hcomp 
557           lapply (Htrue ls c x1 ci tl1 ls0 (cj::tl2') ???)
558           [ #c0 #Hc0 cases (orb_true_l … Hc0) #Hc0
559             [ @Hnotend >(\P Hc0) @memb_hd
560             | @Hnotendx1 // ]
561           | >Hmid_dst >Hrs0 %
562           | >Hxs %
563           | * * #_ #_ -Htrue #Htrue lapply (Htrue ?? (refl ??) ?) [  @(Hcomp ?? (refl ??)) ]
564             * #Htb >Htb #Hendci >Hrs0 >Hxs
565             cases (IH ls c xs end rs ? Hnotend Hend Hnotstart) -IH 
566             [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ]
567             #_ #IH lapply Hxs lapply Hnotendx1 -Hxs -Hnotendx1 cases x1 in Hrs0;
568             [ #Hrs0 #_ whd in ⊢ (???%→?); #Hxs
569               cases (IH Hstart (c::ls0) cj tl2' ?)
570               [ -IH * #l * #l1 * #Hll1 #IH % %{(c::l)} %{l1} 
571                 % [ @eq_f @Hll1 ]
572                 #cj0 #l2 #Hcj0 >(IH … Hcj0) >Htb
573                 >change_vec_commute // >change_vec_change_vec
574                 >change_vec_commute [|@sym_not_eq // ] @eq_f3 //
575                 >reverse_cons >associative_append %
576               | #IH %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%]
577                 @not_sub_list_merge
578                 [ #l2 cut (∃xs'.xs = ci::xs')
579                   [ cases xs in Hxs;
580                     [ normalize #H destruct (H) >Hend in Hendci; #H destruct (H)
581                     | #ci' #xs' normalize #H lapply (cons_injective_l ????? H)
582                       #H1 >H1 %{xs'} % ]
583                   ]
584                   * #xs' #Hxs' >Hxs' normalize % #H destruct (H)
585                   lapply (Hcomp … (refl ??)) * /2/
586                 |#t #l2 #l3 % normalize #H lapply (cons_injective_r ????? H)
587                  -H #H >H in IH; #IH cases (IH l2 l3) -IH #IH @IH % ] 
588               | >Htb >nth_change_vec // >Hmid_dst >Hrs0 % ]
589             | #x2 #xs2 normalize in ⊢ (%→?); #Hrs0 #Hnotendxs2 normalize in ⊢ (%→?);
590               #Hxs cases (IH Hstart (c::ls0) x2 (xs2@cj::tl2') ?)
591               [ -IH * #l * #l1 * #Hll1 #IH % %{(c::l)} %{l1} 
592                 % [ @eq_f @Hll1 ]
593                 #cj0 #l2 #Hcj0 >(IH … Hcj0) >Htb
594                 >change_vec_commute // >change_vec_change_vec
595                 >change_vec_commute [|@sym_not_eq // ] @eq_f3 //
596                 >reverse_cons >associative_append %
597               | -IH #IH %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%] 
598                 @not_sub_list_merge_2 [| @IH]
599                 cut (∃l2.xs = (x2::xs2)@ci::l2)
600                 [lapply Hnotendxs2
601                  lapply Hnotend -Hnotend lapply Hxs
602                  >(?:x2::xs2@ci::tl1 = (x2::xs2)@ci::tl1) [|%]
603                  lapply (x2::xs2) elim xs
604                   [ *
605                     [ normalize in ⊢ (%→?); #H1 destruct (H1) 
606                       >Hendci in Hend; #Hend destruct (Hend)
607                     | #x3 #xs3 normalize in ⊢ (%→?); #H1 destruct (H1)
608                       #_ #Hnotendx3 >(Hnotendx3 ? (memb_hd …)) in Hend;
609                       #Hend destruct (Hend)
610                     ]
611                   | #x3 #xs3 #IHin *
612                     [ normalize in ⊢ (%→?); #Hxs3 destruct (Hxs3) #_ #_
613                       %{xs3} %
614                     | #x4 #xs4 normalize in ⊢ (%→?); #Hxs3xs4 #Hnotend
615                       #Hnotendxs4 destruct (Hxs3xs4) cases (IHin ? e0 ??)
616                       [ #l0 #Hxs3 >Hxs3 %{l0} %
617                       | #c0 #Hc0 @Hnotend cases (orb_true_l … Hc0) -Hc0 #Hc0
618                         [ >(\P Hc0) @memb_hd
619                         | @memb_cons @memb_cons @Hc0 ]
620                       | #c0 #Hc0 @Hnotendxs4 @memb_cons //
621                       ]
622                     ]
623                   ]
624                 ] * #l2 #Hxs'
625                 >Hxs' #l3 normalize >associative_append normalize % #H
626                 destruct (H) lapply (append_l2_injective ?????? e1) //
627                 #H1 destruct (H1) cases (Hcomp ?? (refl ??)) /2/
628               | >Htb >nth_change_vec // >Hmid_dst >Hrs0 % ]
629             ]
630           ]
631         ]
632       ]
633     |lapply (\Pf eqx) -eqx #eqx >Hmid_dst #Htrue 
634      cases (Htrue ? (refl ??) eqx) -Htrue #Htb #Hendcx #_
635      cases rs0 in Htb;
636      [ #_ %2 #l #l1 cases l
637        [ normalize cases xs
638          [ cases l1
639            [ normalize % #H destruct (H) cases eqx /2/
640            | #tmp1 #l2 normalize % #H destruct (H) ]
641          | #tmp1 #l2 normalize % #H destruct (H) ]
642        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
643          [ normalize #H1 destruct (H1)
644          | #tmp2 #l3 normalize #H1 destruct (H1) ]
645        ]
646      | #r1 #rs1 normalize in ⊢ (???(????%?)→?); #Htb >Htb in IH; #IH
647        cases (IH ls x xs end rs ? Hnotend Hend Hnotstart) 
648        [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ] -IH
649        #_ #IH cases (IH Hstart (c::ls0) r1 rs1 ?)
650        [|| >nth_change_vec // ] -IH
651        [ * #l * #l1 * #Hll1 #Hout % %{(c::l)} %{l1} % >Hll1 //
652          >reverse_cons >associative_append #cj0 #ls #Hl1 >(Hout ?? Hl1)
653          >change_vec_commute in ⊢ (??(???%??)?); // @sym_not_eq //
654        | #IH %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@IH]
655          #l1 % #H destruct (H) cases eqx /2/
656        ] 
657     ]
658   ]
659 ]
660 qed.
661
662 definition Pre_match_m ≝ 
663   λsrc,sig,n,is_startc,is_endc.λt: Vector (tape sig) (S n).
664   ∃start,xs,end.
665   nth src (tape sig) t (niltape sig) = midtape ? [] start (xs@[end]) ∧
666   is_startc start = true ∧
667   (∀c.c ∈ (xs@[end]) = true → is_startc c = false) ∧
668   (∀c.c ∈ (start::xs) = true → is_endc c = false) ∧
669   is_endc end = true.
670   
671 lemma terminate_match_m :
672   ∀src,dst,sig,n,is_startc,is_endc,t.
673   src ≠ dst → src < S n → dst < S n → 
674   Pre_match_m src sig n is_startc is_endc t → 
675   match_m src dst sig n is_startc is_endc ↓ t.
676 #src #dst #sig #n #is_startc #is_endc #t #Hneq #Hsrc #Hdst * #start * #xs * #end
677 * * * * #Hmid_src #Hstart #Hnotstart #Hnotend #Hend
678 @(terminate_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst)) //
679 <(change_vec_same … t dst (niltape ?))
680 lapply (refl ? (nth dst (tape sig) t (niltape ?))) 
681 cases (nth dst (tape sig) t (niltape ?)) in ⊢ (???%→?);
682 [ #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
683   >Hmid_src #HR cases (HR ? (refl ??)) -HR
684   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
685   * #H @False_ind @H %
686 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
687   >Hmid_src #HR cases (HR ? (refl ??)) -HR
688   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
689   * #H @False_ind @H %
690 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
691   >Hmid_src #HR cases (HR ? (refl ??)) -HR
692   >nth_change_vec // >Htape_dst normalize in ⊢ (%→?);
693   * #H @False_ind @H %
694 | #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs
695   [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
696    >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ? (refl ??)) -HR #_
697    #HR cases (HR Hstart Hnotstart)
698    cases (true_or_false (start == s)) #Hs
699    [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
700      cut (∃ci,xs1.xs@[end] = ci::xs1)
701      [ cases xs
702        [ %{end} %{[]} %
703        | #x1 #xs1 %{x1} %{(xs1@[end])} % ] ] * #ci * #xs1 #Hxs
704      >Hxs in Htrue; #Htrue
705      cases (Htrue [ ] start [ ] ? xs1 ? [ ] (refl ??) (refl ??) ?)
706      [ * #_ * #H @False_ind @H % ]
707      #c0 #Hc0 @Hnotend >(memb_single … Hc0) @memb_hd
708    | lapply (\Pf Hs) -Hs #Hs #Htrue #_
709      cases (Htrue ? (refl ??) Hs) -Htrue #Ht1 #_ %
710      #t2 whd in ⊢ (%→?); #HR cases (HR start ?)
711      [ >Ht1 >nth_change_vec // normalize in ⊢ (%→?); * #H @False_ind @H %
712      | >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
713        >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src % ]
714    ]
715   |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?);
716    >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src
717    #Htrue cases (Htrue ? (refl ??)) -Htrue #_ #Htrue
718    <(change_vec_same … t1 dst (niltape ?))
719    cases (Htrue Hstart Hnotstart) -Htrue
720    cases (true_or_false (start == s)) #Hs
721    [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
722     cut (∃ls0,xs0,ci,rs,rs0.
723       nth src ? t (niltape ?) = midtape sig [ ] start (xs0@ci::rs) ∧
724       nth dst ? t (niltape ?) = midtape sig ls0 s (xs0@rs0) ∧
725       (is_endc ci = true ∨ (is_endc ci = false ∧ (∀b,tlb.rs0 = b::tlb → ci ≠ b))))
726     [cases (comp_list ? (xs@[end]) (r0::rs0) is_endc) #xs0 * #xs1 * #xs2
727       * * * #Hxs #Hrs #Hxs0notend #Hcomp >Hrs
728       cut (∃y,ys. xs1 = y::ys)
729       [ lapply Hxs0notend lapply Hxs lapply xs0 elim xs
730         [ *
731           [ normalize #Hxs1 <Hxs1 #_ %{end} %{[]} %
732           | #z #zs normalize in ⊢ (%→?); #H destruct (H) #H
733             lapply (H ? (memb_hd …)) -H >Hend #H1 destruct (H1)
734           ]
735         | #y #ys #IH0 * 
736           [ normalize in ⊢ (%→?); #Hxs1 <Hxs1 #_ %{y} %{(ys@[end])} %
737           | #z #zs normalize in ⊢ (%→?); #H destruct (H) #Hmemb
738             @(IH0 ? e0 ?) #c #Hc @Hmemb @memb_cons // ] ] ] * #y * #ys #Hxs1
739       >Hxs1 in Hxs; #Hxs >Hmid_src >Hmid_dst >Hxs >Hrs
740       %{ls} %{xs0} %{y} %{ys} %{xs2}
741       % [ % // | @Hcomp // ] ]
742     * #ls0 * #xs0 * #ci * #rs * #rs0 * * #Hmid_src' #Hmid_dst' #Hcomp
743     <Hmid_src in Htrue; >nth_change_vec // >Hs #Htrue destruct (Hs)
744     lapply (Htrue ??????? Hmid_src' Hmid_dst' ?) -Htrue
745     [ #c0 #Hc0 @Hnotend cases (orb_true_l … Hc0) -Hc0 #Hc0
746       [ whd in ⊢ (??%?); >Hc0 %
747       | @memb_cons >Hmid_src in Hmid_src'; #Hmid_src' destruct (Hmid_src')
748         lapply e0 -e0 @(list_elim_left … rs)
749         [ #e0 destruct (e0) lapply (append_l1_injective_r ?????? e0) //
750         | #x1 #xs1 #_ >append_cons in ⊢ (???%→?);
751           <associative_append #e0 lapply (append_l1_injective_r ?????? e0) //
752           #e1 >e1 @memb_append_l1 @memb_append_l1 // ] ]
753     | * * #Hciendc cases rs0 in Hcomp;
754       [ #_ * #H @False_ind @H %
755       | #r1 #rs1 * [ >Hciendc #H destruct (H) ]
756         * #_ #Hcomp lapply (Hcomp ?? (refl ??)) -Hcomp #Hcomp #_ #Htrue
757         cases (Htrue ?? (refl ??) Hcomp) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
758         [ >nth_change_vec_neq [|@sym_not_eq //]
759           >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
760         | >nth_change_vec // >Hmid_dst % ] ] ]
761   | >Hmid_dst >nth_change_vec // lapply (\Pf Hs) -Hs #Hs #Htrue #_
762     cases (Htrue ? (refl ??) Hs) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
763     [ >nth_change_vec_neq [|@sym_not_eq //]
764       >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
765     | >nth_change_vec // ] ] ] ]
766 qed.