]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/match.ma
progress
[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_strong ≝ λ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   (∀x,x0,xs,rs.
124     nth dst ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
125     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
126     nth src ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
127     outt = change_vec ?? 
128            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) dst)
129            (midtape sig ls0 y0 (reverse ? target@y::rs0)) src) ∧
130   (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → 
131    ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → 
132    outt = int).
133
134 definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
135   (∀x,x0,xs,rs.
136     nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
137     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
138     nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
139     outt = change_vec ?? 
140            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
141            (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧
142   (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → 
143    ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → 
144    outt = int).
145
146 (*
147 theorem accRealize_to_Realize :
148   ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
149   M ⊨ [ acc: Rtrue, Rfalse ] →  M ⊨ Rtrue ∪ Rfalse.
150 #sig #n #M #Rtrue #Rfalse #acc #HR #t
151 cases (HR t) #k * #outc * * #Hloop
152 #Htrue #Hfalse %{k} %{outc} % // 
153 cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
154 [ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
155 qed.
156 *)
157
158 lemma sem_rewind_strong : ∀src,dst,sig,n.
159   src ≠ dst → src < S n → dst < S n → 
160   rewind src dst sig n ⊨ R_rewind_strong src dst sig n.
161 #src #dst #sig #n #Hneq #Hsrc #Hdst
162 @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst) ?)
163 [| @(sem_seq_app sig n ????? (sem_move_multi … R ?) (sem_move_multi … R ?)) //
164  @le_S_S_to_le // ]
165 #ta #tb * #tc * * * #Htc1 #Htc2 #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb % [ %
166 [ #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst
167   >(Htc1 ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd;
168   [|>Hmidta_dst //
169   |>length_append >length_append >Hlen % ]
170   >change_vec_commute [|@sym_not_eq //]
171   >change_vec_change_vec
172   >nth_change_vec_neq [|@sym_not_eq //]
173   >nth_change_vec // >reverse_append >reverse_single
174   >reverse_append >reverse_single normalize in match (tape_move ???);
175   >rev_append_def >append_nil #Htd >Htd in Htb;
176   >change_vec_change_vec >nth_change_vec //
177   cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); //
178 | #x #x0 #xs #rs #Hmidta_dst #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_src
179   >(Htc2 ??? Hmidta_dst ls0 y (target@[y0]) rs0 ??) in Htd;
180   [|>Hmidta_src //
181   |>length_append >length_append >Hlen % ]
182   >change_vec_change_vec
183   >change_vec_commute [|@sym_not_eq //]
184   >nth_change_vec // 
185   >reverse_append >reverse_single
186   >reverse_append >reverse_single
187   cases ls0 [|#l1 #ls1] normalize in match (tape_move ???);
188   #Htd >Htd in Htb; >change_vec_change_vec >nth_change_vec //
189   >rev_append_def >change_vec_commute // normalize in match (tape_move ???); // ]
190 | #x #rs #Hmidta_src #ls0 #y #rs0 #Hmidta_dst
191   lapply (Htc1 … Hmidta_src … (refl ??) Hmidta_dst) -Htc1 #Htc >Htc in Htd;
192   >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
193   >nth_change_vec_neq [|@sym_not_eq //] 
194   >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%);
195   [ #Hls0 #Htd >Htd in Htb; 
196     >nth_change_vec // >change_vec_change_vec 
197     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
198     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
199   | #l1 #ls1 #Hls0 #Htd >Htd in Htb;
200     >nth_change_vec // >change_vec_change_vec 
201     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
202     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
203 ]]
204 qed.
205
206 lemma sem_rewind : ∀src,dst,sig,n.
207   src ≠ dst → src < S n → dst < S n → 
208   rewind src dst sig n ⊨ R_rewind src dst sig n.
209 #src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_rewind_strong …)) //
210 #ta #tb * * #H1 #H2 #H3 % /2/
211 qed.
212
213 definition match_step ≝ λsrc,dst,sig,n.
214   compare src dst sig n ·
215      (ifTM ?? (partest sig n (match_test src dst sig ?))
216       (single_finalTM ??
217         (rewind src dst sig n · (inject_TM ? (move_r ?) n dst)))
218       (nop …)
219       partest1).
220
221 (* we assume the src is a midtape
222    we stop
223    if the dst is out of bounds (outt = int)
224    or dst.right is shorter than src.right (outt.current → None)
225    or src.right is a prefix of dst.right (out = just right of the common prefix) *)
226 definition R_match_step_false ≝ 
227   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
228   ∀ls,x,xs.
229   nth src ? int (niltape ?) = midtape sig ls x xs →
230   ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
231     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
232       xs = rs0@xs0 ∧
233       outt = change_vec ??
234              (change_vec ?? int (mk_tape sig (reverse ? rs0@x::ls) (option_hd ? xs0) (tail ? xs0)) src)
235              (mk_tape ? (reverse ? rs0@x::ls0) (None ?) [ ]) dst) ∨
236     (∃ls0,rs0. 
237      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
238      (* ∀rsj,c. 
239      rs0 = c::rsj → *)
240      outt = change_vec ??
241             (change_vec ?? int (mk_tape sig (reverse ? xs@x::ls) (None ?) [ ]) src)
242             (mk_tape sig (reverse ? xs@x::ls0) (option_hd ? rs0) (tail ? rs0)) dst).
243
244 (*
245    we assume the src is a midtape [ ] s rs
246    if we iterate
247    then dst.current = Some ? s1
248    and  if s ≠ s1 then outt = int.dst.move_right()
249    and  if s = s1 
250         then int.src.right and int.dst.right have a common prefix
251         and  the heads of their suffixes are different
252         and  outt = int.dst.move_right().
253                
254  *)
255 definition R_match_step_true ≝ 
256   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
257   ∀s,rs.nth src ? int (niltape ?) = midtape ? [ ] s rs → 
258   outt = change_vec ?? int 
259          (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst ∧
260   (∃s0.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s0 ∧
261    (s0 = s →  
262    ∃xs,ci,rs',ls0,cj,rs0.
263    rs = xs@ci::rs' ∧
264    nth dst ? int (niltape ?) = midtape sig ls0 s (xs@cj::rs0) ∧
265    ci ≠ cj)).
266   
267 lemma sem_match_step :
268   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
269   match_step src dst sig n ⊨ 
270     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
271       R_match_step_true src dst sig n, 
272       R_match_step_false src dst sig n ].
273 #src #dst #sig #n #Hneq #Hsrc #Hdst 
274 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
275     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
276       (sem_seq … 
277         (sem_rewind ???? Hneq Hsrc Hdst) 
278         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
279       (sem_nop …)))
280 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
281   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
282   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
283     whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
284     >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
285   | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
286     cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
287     [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
288       whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
289       >nth_current_chars >nth_current_chars >Hs >Hcurta_dst 
290       normalize in ⊢ (%→?); #H destruct (H)
291     | #s0 #Hs0
292       cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
293       cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
294       cases (true_or_false (s == s0)) #Hss0
295       [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0) 
296         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
297         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
298           >nth_current_chars >nth_current_chars >Hcurtc_dst 
299           cases (current ? (nth src …)) 
300           [normalize in ⊢ (%→?); #H destruct (H)
301           | #x >nth_change_vec // cases (reverse ? rs0)
302             [ normalize in ⊢ (%→?); #H destruct (H)
303             | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
304         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
305           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
306           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
307             [>nth_change_vec [cases (append ???) // | @Hsrc] 
308             |@(not_to_not … Hneq) //
309             ]]
310           normalize in ⊢ (%→?); #H destruct (H) ]
311         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
312           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * *
313           >Htc >change_vec_commute // >nth_change_vec //
314           >change_vec_commute [|@sym_not_eq //] >nth_change_vec // #Hte #_ #Htb
315           #s' #rs' >Hmidta_src #H destruct (H)
316           lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
317           >change_vec_commute // >change_vec_change_vec
318           >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
319           >Hte in Htb; * * #_ >nth_change_vec // #Htb1
320           lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 %
321           [ @(eq_vec … (niltape ?)) #i #Hi
322             cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
323             [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst
324               >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] %
325             | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)]
326               >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%); 
327               <Hrs <Hmidta_src [|@(\Pf Hdsti)] >change_vec_same % ]
328           | >Hmidta_dst %{s'} % [%] #_
329             >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
330           ]
331         ]
332       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
333         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
334         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
335         #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2
336         #s1 #rs1 >Hmidta_src #H destruct (H)
337         lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
338         [ @(eq_vec … (niltape ?)) #i #Hi
339           cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
340           [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
341             cases rs0 [|#r2 #rs2] %
342           | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
343         | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ]
344       ]
345     ]
346   ]
347 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
348   whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
349   lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
350   cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
351   [ #Hcurta_dst % % % // @Hcomp1 %2 //
352   | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
353     #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
354     [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
355     | >(?:tc=ta) in Htest; 
356       [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
357         #Hxx0' destruct (Hxx0') % ]
358       whd in ⊢ (??%?→?); 
359       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
360       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
361       cases (Hcomp2 … Hta_src Hta_dst) [ *
362       [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
363         [ % // | >Hcurtc % ]
364       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
365       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
366         #Hci #Hxs #Hrs0 #Htc @False_ind
367         whd in Htest:(??%?); 
368         >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest; 
369         [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
370           >nth_change_vec //]
371         >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj) 
372         [|>nth_current_chars >Htc >nth_change_vec //]
373         normalize #H destruct (H) ] ] ]
374 qed.
375
376 definition match_m ≝ λsrc,dst,sig,n.
377   whileTM … (match_step src dst sig n) 
378     (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
379
380 definition R_match_m ≝ 
381   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
382   ∀x,rs.
383   nth src ? int (niltape ?) = midtape sig [ ] x rs →
384   (current sig (nth dst (tape sig) int (niltape sig)) = None ? → 
385    right ? (nth dst (tape sig) int (niltape sig)) = [ ] → outt = int) ∧
386   (∀ls0,x0,rs0.
387    nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
388    (∃l,l1.x0::rs0 = l@x::rs@l1 ∧
389     outt = change_vec ?? 
390            (change_vec ?? int 
391              (mk_tape sig (reverse ? rs@[x]) (None ?) [ ]) src)
392            (mk_tape sig ((reverse ? (l@x::rs))@ls0) (option_hd ? l1) (tail ? l1)) dst) ∨
393     ∀l,l1.x0::rs0 ≠ l@x::rs@l1).
394
395 lemma not_sub_list_merge : 
396   ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
397 #T #a #b #H1 #H2 #l elim l normalize //
398 qed.
399
400 lemma not_sub_list_merge_2 : 
401   ∀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.
402 #T #a #b #t #H1 #H2 #l elim l //
403 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
404 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
405 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
406 qed.
407
408
409 lemma wsem_match_m : ∀src,dst,sig,n.
410 src ≠ dst → src < S n → dst < S n → 
411   match_m src dst sig n ⊫ R_match_m src dst sig n.
412 #src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
413 lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
414 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
415 [ #Hfalse #x #xs #Hmid_src
416   cases (Hfalse … Hmid_src) -Hfalse 
417   [(* current dest = None *) *
418     [ * #Hcur_dst #Houtc %
419       [#_ >Houtc //
420       | #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst; 
421         normalize in ⊢ (%→?); #H destruct (H)
422       ]
423     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
424       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
425       | #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
426         >Hrs0 >HNone cases xs0
427         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
428           @eq_f3 //
429           [ >reverse_append %
430           | >reverse_append >reverse_cons >reverse_append
431             >associative_append >associative_append % ]
432         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
433           >length_append whd in ⊢ (??%(??%)→?); >length_append
434           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
435           #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
436           >associative_plus >associative_plus 
437           #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
438           #e2 destruct (e2)
439         ]
440       ]
441     ]
442   |* #ls0 * #rs0 * #Hmid_dst #Houtc %
443     [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
444     |#ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
445      %1 %{[ ]} %{rs0} % [%] 
446      >reverse_cons >associative_append >Houtc %
447     ]
448   ]
449 |-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
450  #x #xs #Hmidta_src
451  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
452  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
453   [#Hcurta_dst % 
454     [#Hcurta_dst #Hrightta_dst whd in Htrue; >Hmidta_src in Htrue; #Htrue
455      cases (Htrue ?? (refl ??)) -Htrue #Htc
456      cut (tc = ta)
457      [ >Htc whd in match (tape_move_mono ???); whd in match (tape_write ???);
458        <(change_vec_same … ta dst (niltape ?)) in ⊢ (???%);
459        lapply Hrightta_dst lapply Hcurta_dst -Hrightta_dst -Hcurta_dst 
460        cases (nth dst ? ta (niltape ?))
461        [ #_ #_ %
462        | #r0 #rs0 #_ normalize in ⊢ (%→?); #H destruct (H)
463        | #l0 #ls0 #_ #_ %
464        | #ls #x0 #rs normalize in ⊢ (%→?); #H destruct (H) ] ]
465      -Htc #Htc destruct (Htc) #_
466      cases (IH … Hmidta_src) #Houtc #_ @Houtc //
467     |#ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst;
468      normalize in ⊢ (%→?); #H destruct (H)
469     ]
470   | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
471     #ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst; normalize in ⊢ (%→?);
472     #H destruct (H) whd in Htrue; >Hmidta_src in Htrue; #Htrue
473     cases (Htrue ?? (refl …)) -Htrue >Hmidta_dst #Htc
474     cases (true_or_false (x==c)) #eqx
475     [ lapply (\P eqx) -eqx #eqx destruct (eqx) * #s0 * whd in ⊢ (??%?→?); #Hs0
476       destruct (Hs0) #Htrue cases (Htrue (refl ??)) -Htrue
477       #xs0 * #ci * #rs' * #ls1 * #cj * #rs1 * * #Hxs #H destruct (H) #Hcicj
478       >Htc in IH; whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
479       #IH cases (IH … Hmidta_src) -IH #_ >nth_change_vec //
480       cut (∃x1,xs1.xs0@cj::rs1 = x1::xs1)
481       [ cases xs0 [ %{cj} %{rs1} % | #x1 #xs1 %{x1} %{(xs1@cj::rs1)} % ] ] * #x1 * #xs1
482       #Hxs1 >Hxs1 #IH cases (IH … (refl ??)) -IH
483       [ * #l * #l1 * #Hxs1'
484         >change_vec_commute // >change_vec_change_vec
485         #Houtc % %{(s0::l)} %{l1} % 
486         [ normalize <Hxs1' %
487         | >reverse_cons >associative_append >change_vec_commute // @Houtc ]
488       | #H %2 #l #l1 >(?:l@s0::xs@l1 = l@(s0::xs)@l1) [|%]
489         @not_sub_list_merge
490         [ #l2 >Hxs <Hxs1 % normalize #H1 lapply (cons_injective_r ????? H1)
491           >associative_append #H2 lapply (append_l2_injective ????? (refl ??) H2)
492           #H3 lapply (cons_injective_l ????? H3) #H3 >H3 in Hcicj; * /2/
493         |#t #l2 #l3 % normalize #H1 lapply (cons_injective_r ????? H1)
494          -H1 #H1 cases (H l2 l3) #H2 @H2 @H1
495         ]
496       ]
497     | #_ cases (IH x xs ?) -IH
498       [| >Htc >nth_change_vec_neq [|@sym_not_eq //] @Hmidta_src ]
499       >Htc >nth_change_vec // cases rs0
500       [ #_ #_ %2 #l #l1 cases l
501        [ normalize cases xs
502          [ cases l1
503            [ normalize % #H destruct (H) cases (\Pf eqx) /2/
504            | #tmp1 #l2 normalize % #H destruct (H) ]
505          | #tmp1 #l2 normalize % #H destruct (H) ]
506        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
507          [ normalize #H1 destruct (H1)
508          | #tmp2 #l3 normalize #H1 destruct (H1) ] ]
509       | #r1 #rs1 #_ #IH cases (IH … (refl ??)) -IH
510         [ * #l * #l1 * #Hll1 #Houtc % %{(c::l)} %{l1} % [ >Hll1 % ]
511           >Houtc >change_vec_commute // >change_vec_change_vec 
512           >change_vec_commute [|@sym_not_eq //]
513           >reverse_cons >associative_append %
514         | #Hll1 %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@Hll1]
515          #l1 % #H destruct (H) cases (\Pf eqx) /2/
516         ]
517       ]
518     ]
519   ]
520 ]
521 qed.
522
523 axiom daemon : ∀P:Prop.P.
524
525 definition R_match_step_true_naive ≝ 
526   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
527   |left ? (nth src ? outt (niltape ?))| +
528   |right ? (nth dst ? outt (niltape ?))| <
529   |left ? (nth src ? int (niltape ?))| +
530   |right ? (nth dst ? int (niltape ?))|.
531
532 axiom right_mk_tape : ∀sig,ls,c,rs.right ? (mk_tape sig ls c rs) = rs.
533 axiom left_mk_tape : ∀sig,ls,c,rs.left ? (mk_tape sig ls c rs) = ls.
534 axiom length_tail : ∀A,l.0 < |l| → |tail A l| < |l|.
535 axiom lists_length_split : 
536  ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
537   
538 lemma sem_match_step_termination :
539   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
540   match_step src dst sig n ⊨ 
541     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
542       R_match_step_true_naive src dst sig n, 
543       R_match_step_false src dst sig n ].
544 #src #dst #sig #n #Hneq #Hsrc #Hdst 
545 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
546     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
547       (sem_seq … 
548         (sem_rewind_strong ???? Hneq Hsrc Hdst) 
549         (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
550       (sem_nop …)))
551 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
552   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
553   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
554     whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
555     >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
556   | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
557     cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
558     [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
559       whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
560       >nth_current_chars >nth_current_chars >Hs >Hcurta_dst 
561       normalize in ⊢ (%→?); #H destruct (H)
562     | #s0 #Hs0
563       cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
564       cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
565       cases (true_or_false (s == s0)) #Hss0
566       [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0) 
567         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
568         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
569           >nth_current_chars >nth_current_chars >Hcurtc_dst 
570           cases (current ? (nth src …)) 
571           [normalize in ⊢ (%→?); #H destruct (H)
572           | #x >nth_change_vec // cases (reverse ? rs0)
573             [ normalize in ⊢ (%→?); #H destruct (H)
574             | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
575         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
576           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
577           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
578             [>nth_change_vec [cases (append ???) // | @Hsrc] 
579             |@(not_to_not … Hneq) //
580             ]]
581           normalize in ⊢ (%→?); #H destruct (H) ]
582         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
583           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * *
584           >Htc >change_vec_commute // >nth_change_vec //
585           >change_vec_commute [|@sym_not_eq //] >nth_change_vec //
586           cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb
587           destruct (Hlsalsb)  *
588           [ #Hte #_ #_ <(reverse_reverse … ls) in Hte; <(reverse_reverse … lsa)
589             cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen' 
590             @(list_cases2 … Hlen')
591             [ #H1 #H2 >H1 >H2 -H1 -H2 normalize in match (reverse ? [ ]); #Hte
592               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
593               >change_vec_commute // >change_vec_change_vec
594               >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
595               >Hte * * #_ >nth_change_vec // >reverse_reverse 
596               #H lapply (H … (refl ??)) -H #Htb1 #Htb2
597               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (midtape sig [] s0 (xs@ci::rs'')) src) (mk_tape sig (s0::lsb) (option_hd sig (xs@cj::rs0')) (tail sig (xs@cj::rs0'))) dst)
598               [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
599               >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
600               >right_mk_tape normalize in match (left ??);
601               >Hmidta_src >Hmidta_dst >length_tail >Hrs0 >length_append 
602               normalize in ⊢ (?(?%)(?%?)); >commutative_plus //
603             | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
604               >reverse_cons >reverse_cons #Hte
605               lapply (Hte ci hdb (reverse ? xs@s0::reverse ? tlb) rs'' ?
606                        lsb cj hda (reverse ? xs@s0::reverse ? tla) rs0' ??)
607               [ /2 by cons_injective_l, nil/
608               | >length_append >length_append @eq_f @(eq_f ?? S)
609                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
610                 >length_reverse >length_reverse destruct (Hlen') //
611               | /2 by refl, trans_eq/ ] -Hte
612               #Hte * * #_ >Hte >nth_change_vec // #Htb1 #Htb2
613               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
614                         (mk_tape sig (hda::lsb) (option_hd ? (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0')) (tail ? (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0'))) dst) 
615                         (midtape ? [ ] hdb (reverse sig (reverse sig xs@s0::reverse sig tlb)@ci::rs'')) src)
616               [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
617               >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
618               >right_mk_tape >Hmidta_src >Hmidta_dst 
619               whd in match (left ??); whd in match (left ??); whd in match (right ??);
620               >length_tail >Hrs0 >length_append >length_append >length_reverse
621               >length_append >commutative_plus in match (|reverse ??| + ?);
622               whd in match (|?::?|); >length_reverse >length_reverse
623               <(length_reverse ? ls) <Hlen' >H1 normalize // ]
624          | #_ #Hte #_ <(reverse_reverse … ls0) in Hte; <(reverse_reverse … lsa)
625             cut (|reverse ? lsa| = |reverse ? ls0|) [ // ] #Hlen' 
626             @(list_cases2 … Hlen')
627             [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #Hte
628               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
629               >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] 
630               >change_vec_change_vec #Hte
631               >Hte * * #_ >nth_change_vec // >reverse_reverse 
632               #H lapply (H … (refl ??)) -H #Htb1 #Htb2
633               (* cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (midtape sig [] s0 (xs@ci::rs'')) src) (mk_tape sig (s0::lsb) (option_hd sig (xs@cj::rs0')) (tail sig (xs@cj::rs0'))) dst) *)
634               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (mk_tape ? [s0] (option_hd ? (xs@cj::rs0')) (tail ? (xs@cj::rs0'))) dst)
635                            (midtape ? lsb s0 (xs@ci::rs'')) src)
636               [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
637               >nth_change_vec_neq // >nth_change_vec //
638               >right_mk_tape normalize in match (left ??);
639               >Hmidta_src >Hmidta_dst >length_tail >Hrs0 >length_append 
640               >commutative_plus in match (pred ?); normalize
641               >length_append >(?:|lsa| = O)
642               [ normalize <plus_n_Sm <plus_n_Sm // | <(length_reverse ? lsa) >H1 % ]
643             | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
644               >reverse_cons >reverse_cons #Hte
645               lapply (Hte cj hdb (reverse ? xs@s0::reverse ? tlb) rs0' ?
646                        lsb ci hda (reverse ? xs@s0::reverse ? tla) rs'' ??)
647               [ /2 by cons_injective_l, nil/
648               | >length_append >length_append @eq_f @(eq_f ?? S)
649                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
650                 >length_reverse >length_reverse destruct (Hlen') //
651               | /2 by refl, trans_eq/ ] -Hte
652               #Hte * * #_ >Hte >nth_change_vec_neq // >nth_change_vec // #Htb1 #Htb2
653               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
654                         (mk_tape sig [hdb] (option_hd ? (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0')) (tail ? (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0'))) dst) 
655                         (midtape ? lsb hda (reverse sig (reverse sig xs@s0::reverse sig tla)@ci::rs'')) src)
656               [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
657               >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
658               >right_mk_tape >Hmidta_src >Hmidta_dst 
659               whd in match (left ??); whd in match (left ??); whd in match (right ??);
660               >length_tail >Hrs0 >length_append >length_append >length_reverse
661               >length_append >commutative_plus in match (|reverse ??| + ?);
662               whd in match (|?::?|); >length_reverse >length_reverse >Hlen
663               <(length_reverse ? ls0) >H2 whd in match (|?::?|); 
664               >length_append normalize //
665             ]
666           ]
667         ]
668       | FIXME
669        (* lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
670         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
671         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
672         #Htd destruct (Htd) * #te * * * # #_ >Hmidta_src >Hmidta_dst #Hte
673         lapply (Hte … (refl ??) … (refl ??)) * * #_ #Htb1 #Htb2
674         #s1 #rs1 >Hmidta_src #H destruct (H)
675         lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
676         [ @(eq_vec … (niltape ?)) #i #Hi
677           cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
678           [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
679             cases rs0 [|#r2 #rs2] %
680           | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
681         | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ] *)
682       ]
683     ]
684   ]
685 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
686   whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
687   lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
688   cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
689   [ #Hcurta_dst % % % // @Hcomp1 %2 //
690   | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
691     #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
692     [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
693     | >(?:tc=ta) in Htest; 
694       [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
695         #Hxx0' destruct (Hxx0') % ]
696       whd in ⊢ (??%?→?); 
697       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
698       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
699       cases (Hcomp2 … Hta_src Hta_dst) [ *
700       [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
701         [ % // | >Hcurtc % ]
702       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
703       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
704         #Hci #Hxs #Hrs0 #Htc @False_ind
705         whd in Htest:(??%?); 
706         >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest; 
707         [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
708           >nth_change_vec //]
709         >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj) 
710         [|>nth_current_chars >Htc >nth_change_vec //]
711         normalize #H destruct (H) ] ] ]
712 qed.
713
714
715 definition Pre_match_m ≝ 
716   λsrc,sig,n.λt: Vector (tape sig) (S n).
717   ∃x,xs.
718   nth src (tape sig) t (niltape sig) = midtape ? [] x xs.
719   
720 lemma terminate_match_m :
721   ∀src,dst,sig,n,t.
722   src ≠ dst → src < S n → dst < S n → 
723   Pre_match_m src sig n t → 
724   match_m src dst sig n ↓ t.
725 #src #dst #sig #n #t #Hneq #Hsrc #Hdst * #start * #xs
726 #Hmid_src
727 @(terminate_while … (sem_match_step src dst sig n Hneq Hsrc Hdst)) //
728 <(change_vec_same … t dst (niltape ?))
729 lapply (refl ? (nth dst (tape sig) t (niltape ?))) 
730 cases (nth dst (tape sig) t (niltape ?)) in ⊢ (???%→?);
731 [ #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
732   >Hmid_src #HR cases (HR ?? (refl ??)) -HR
733   >nth_change_vec // >Htape_dst #_ * #s0 * normalize in ⊢ (%→?); #H destruct (H)
734 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
735   >Hmid_src #HR cases (HR ?? (refl ??)) -HR
736   >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
737   * #s0 * #H destruct (H)
738 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
739   >Hmid_src #HR cases (HR ?? (refl ??)) -HR
740   >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
741   * #s0 * #H destruct (H)
742 | #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs
743   [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
744    >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ?? (refl ??)) -HR 
745    >change_vec_change_vec #Ht1 #_ % #t2 whd in ⊢ (%→?);
746    >Ht1 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src #HR
747    cases (HR ?? (refl ??)) -HR #_
748    >nth_change_vec // * #s1 * normalize in ⊢ (%→?); #H destruct (H)
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 >change_vec_change_vec
752    >nth_change_vec // >Hmid_dst whd in match (tape_move_mono ???); #Ht1
753    * #s0 * whd in ⊢ (??%?→?); #H destruct (H) #_ >Ht1
754    lapply (IH t1 ? (s0::ls) r0 ?) 
755    [ >Ht1 >nth_change_vec //
756    | >Ht1 >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
757    | >Ht1 >nth_change_vec // ]
758  ]
759 ]
760 qed.