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