]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/match.ma
match now only uses the new move operation
[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/simple_machines.ma".
16 include "turing/multi_universal/compare.ma".
17 include "turing/multi_universal/par_test.ma".
18 include "turing/multi_universal/moves_2.ma".
19
20 lemma eq_vec_change_vec : ∀sig,n.∀v1,v2:Vector sig n.∀i,t,d.
21   nth i ? v2 d = t → 
22   (∀j.i ≠ j → nth j ? v1 d = nth j ? v2 d) → 
23   v2 = change_vec ?? v1 t i.
24 #sig #n #v1 #v2 #i #t #d #H1 #H2 @(eq_vec … d)
25 #i0 #Hlt cases (decidable_eq_nat i0 i) #Hii0
26 [ >Hii0 >nth_change_vec //
27 | >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @H2 @sym_not_eq // ]
28 qed.
29
30 lemma right_mk_tape : 
31   ∀sig,ls,c,rs.(c = None ? → ls = [ ] ∨ rs = [ ]) → right ? (mk_tape sig ls c rs) = rs.
32 #sig #ls #c #rs cases c // cases ls 
33 [ cases rs // 
34 | #l0 #ls0 #H normalize cases (H (refl ??)) #H1 [ destruct (H1) | >H1 % ] ]
35 qed.
36
37 lemma left_mk_tape : ∀sig,ls,c,rs.left ? (mk_tape sig ls c rs) = ls.
38 #sig #ls #c #rs cases c // cases ls // cases rs //
39 qed.
40
41 lemma current_mk_tape : ∀sig,ls,c,rs.current ? (mk_tape sig ls c rs) = c.
42 #sig #ls #c #rs cases c // cases ls // cases rs //
43 qed.
44
45 lemma length_tail : ∀A,l.0 < |l| → |tail A l| < |l|.
46 #A * normalize //
47 qed.
48
49 (*
50 [ ]   → [ ], l2, 1
51 a::al → 
52       [ ]   → [ ], l1, 2
53       b::bl → match rec(al,bl)
54             x, y, 1 → b::x, y, 1
55             x, y, 2 → a::x, y, 2
56 *)
57
58 lemma lists_length_split : 
59  ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
60 #A #l1 elim l1
61 [ #l2 %{[ ]} %{l2} % % %
62 | #hd1 #tl1 #IH *
63   [ %{[ ]} %{(hd1::tl1)} %2 % %
64   | #hd2 #tl2 cases (IH tl2) #x * #y *
65     [ * #IH1 #IH2 %{(hd2::x)} %{y} % normalize % //
66     | * #IH1 #IH2 %{(hd1::x)} %{y} %2 normalize % // ]
67   ]
68 ]
69 qed.
70
71 definition option_cons ≝ λsig.λc:option sig.λl.
72   match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
73
74 lemma opt_cons_tail_expand : ∀A,l.l = option_cons A (option_hd ? l) (tail ? l). 
75 #A * //
76 qed.
77
78 definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
79   match (nth src (option sig) v (None ?)) with 
80   [ None ⇒  false 
81   | Some x ⇒  notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
82   
83 definition rewind ≝ λsrc,dst,sig,n.
84   parmove src dst sig n L · mmove src sig n R · mmove dst sig n R.
85
86 definition R_rewind_strong ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
87   (∀x,x0,xs,rs.
88     nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
89     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
90     nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
91     outt = change_vec ?? 
92            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
93            (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧
94   (∀x,x0,xs,rs.
95     nth dst ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
96     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
97     nth src ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
98     outt = change_vec ?? 
99            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) dst)
100            (midtape sig ls0 y0 (reverse ? target@y::rs0)) src) ∧
101   (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → 
102    ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → 
103    outt = int) ∧
104   (∀x,rs.nth dst ? int (niltape ?) = midtape sig [] x rs → 
105    ∀ls0,y,rs0.nth src ? int (niltape ?) = midtape sig ls0 y rs0 → 
106    outt = int).
107
108 definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
109   (∀x,x0,xs,rs.
110     nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs → 
111     ∀ls0,y,y0,target,rs0.|xs| = |target| → 
112     nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 → 
113     outt = change_vec ?? 
114            (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
115            (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧
116   (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs → 
117    ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 → 
118    outt = int).
119
120 (*
121 theorem accRealize_to_Realize :
122   ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
123   M ⊨ [ acc: Rtrue, Rfalse ] →  M ⊨ Rtrue ∪ Rfalse.
124 #sig #n #M #Rtrue #Rfalse #acc #HR #t
125 cases (HR t) #k * #outc * * #Hloop
126 #Htrue #Hfalse %{k} %{outc} % // 
127 cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
128 [ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
129 qed.
130 *)
131
132 lemma sem_rewind_strong : ∀src,dst,sig,n.
133   src ≠ dst → src < S n → dst < S n → 
134   rewind src dst sig n ⊨ R_rewind_strong src dst sig n.
135 #src #dst #sig #n #Hneq #Hsrc #Hdst
136 @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst) ?)
137 [| @(sem_seq_app sig n ????? (sem_move_multi … R ?) (sem_move_multi … R ?)) //
138  @le_S_S_to_le // ]
139 #ta #tb * #tc * * * #Htc1 #Htc2 #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb % [ % [ %
140 [ #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst
141   >(Htc1 ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd;
142   [|>Hmidta_dst //
143   |>length_append >length_append >Hlen % ]
144   >change_vec_commute [|@sym_not_eq //]
145   >change_vec_change_vec
146   >nth_change_vec_neq [|@sym_not_eq //]
147   >nth_change_vec // >reverse_append >reverse_single
148   >reverse_append >reverse_single normalize in match (tape_move ???);
149   >rev_append_def >append_nil #Htd >Htd in Htb;
150   >change_vec_change_vec >nth_change_vec //
151   cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); //
152 | #x #x0 #xs #rs #Hmidta_dst #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_src
153   >(Htc2 ??? Hmidta_dst ls0 y (target@[y0]) rs0 ??) in Htd;
154   [|>Hmidta_src //
155   |>length_append >length_append >Hlen % ]
156   >change_vec_change_vec
157   >change_vec_commute [|@sym_not_eq //]
158   >nth_change_vec // 
159   >reverse_append >reverse_single
160   >reverse_append >reverse_single
161   cases ls0 [|#l1 #ls1] normalize in match (tape_move ???);
162   #Htd >Htd in Htb; >change_vec_change_vec >nth_change_vec //
163   >rev_append_def >change_vec_commute // normalize in match (tape_move ???); // ]
164 | #x #rs #Hmidta_src #ls0 #y #rs0 #Hmidta_dst
165   lapply (Htc1 … Hmidta_src … (refl ??) Hmidta_dst) -Htc1 #Htc >Htc in Htd;
166   >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
167   >nth_change_vec_neq [|@sym_not_eq //] 
168   >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%);
169   [ #Hls0 #Htd >Htd in Htb; 
170     >nth_change_vec // >change_vec_change_vec 
171     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
172     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
173   | #l1 #ls1 #Hls0 #Htd >Htd in Htb;
174     >nth_change_vec // >change_vec_change_vec 
175     whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
176     <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
177   ] ]
178 | #x #rs #Hmidta_dst #ls0 #y #rs0 #Hmidta_src
179   lapply (Htc2 … Hmidta_dst … (refl ??) Hmidta_src) -Htc2 #Htc >Htc in Htd;
180   >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
181   >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%);
182   [ #Hls0 destruct (Hls0) #Htd >Htd in Htb; 
183     >nth_change_vec // >change_vec_change_vec 
184     whd in match (tape_move ???);whd in match (tape_move ???); 
185     <Hmidta_src <Hmidta_dst >change_vec_same >change_vec_same //
186   | #l1 #ls1 #Hls0 destruct (Hls0) #Htd >Htd in Htb;
187     >nth_change_vec // >change_vec_change_vec 
188     whd in match (tape_move ???); whd in match (tape_move ???); <Hmidta_src
189     <Hmidta_dst >change_vec_same >change_vec_same //
190   ]
191 ]
192 qed.
193
194 lemma sem_rewind : ∀src,dst,sig,n.
195   src ≠ dst → src < S n → dst < S n → 
196   rewind src dst sig n ⊨ R_rewind src dst sig n.
197 #src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_rewind_strong …)) //
198 #ta #tb * * * #H1 #H2 #H3 #H4 % /2 by /
199 qed.
200
201 definition match_step ≝ λsrc,dst,sig,n.
202   compare src dst sig n ·
203      (ifTM ?? (partest sig n (match_test src dst sig ?))
204       (single_finalTM ??
205         (rewind src dst sig n · mmove dst ?? R))
206       (nop …)
207       partest1).
208
209 (* we assume the src is a midtape
210    we stop
211    if the dst is out of bounds (outt = int)
212    or dst.right is shorter than src.right (outt.current → None)
213    or src.right is a prefix of dst.right (out = just right of the common prefix) *)
214 definition R_match_step_false ≝ 
215   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
216   ∀ls,x,xs.
217   nth src ? int (niltape ?) = midtape sig ls x xs →
218   ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
219     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
220       xs = rs0@xs0 ∧
221       outt = change_vec ??
222              (change_vec ?? int (mk_tape sig (reverse ? rs0@x::ls) (option_hd ? xs0) (tail ? xs0)) src)
223              (mk_tape ? (reverse ? rs0@x::ls0) (None ?) [ ]) dst) ∨
224     (∃ls0,rs0. 
225      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
226      (* ∀rsj,c. 
227      rs0 = c::rsj → *)
228      outt = change_vec ??
229             (change_vec ?? int (mk_tape sig (reverse ? xs@x::ls) (None ?) [ ]) src)
230             (mk_tape sig (reverse ? xs@x::ls0) (option_hd ? rs0) (tail ? rs0)) dst).
231
232 (*
233    we assume the src is a midtape [ ] s rs
234    if we iterate
235    then dst.current = Some ? s1
236    and  if s ≠ s1 then outt = int.dst.move_right()
237    and  if s = s1 
238         then int.src.right and int.dst.right have a common prefix
239         and  the heads of their suffixes are different
240         and  outt = int.dst.move_right().
241                
242  *)
243 definition R_match_step_true ≝ 
244   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
245   ∀s,rs.nth src ? int (niltape ?) = midtape ? [ ] s rs → 
246   outt = change_vec ?? int 
247          (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst ∧
248   (∃s0.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s0 ∧
249    (s0 = s →  
250    ∃xs,ci,rs',ls0,cj,rs0.
251    rs = xs@ci::rs' ∧
252    nth dst ? int (niltape ?) = midtape sig ls0 s (xs@cj::rs0) ∧
253    ci ≠ cj)).
254   
255 lemma sem_match_step :
256   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
257   match_step src dst sig n ⊨ 
258     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
259       R_match_step_true src dst sig n, 
260       R_match_step_false src dst sig n ].
261 #src #dst #sig #n #Hneq #Hsrc #Hdst 
262 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
263     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
264       (sem_seq … 
265         (sem_rewind ???? Hneq Hsrc Hdst) 
266         (sem_move_multi … R ?))
267       (sem_nop …))) /2/
268 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
269   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
270   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
271     whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
272     >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
273   | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
274     cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
275     [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
276       whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
277       >nth_current_chars >nth_current_chars >Hs >Hcurta_dst 
278       normalize in ⊢ (%→?); #H destruct (H)
279     | #s0 #Hs0
280       cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
281       cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
282       cases (true_or_false (s == s0)) #Hss0
283       [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0) 
284         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
285         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
286           >nth_current_chars >nth_current_chars >Hcurtc_dst 
287           cases (current ? (nth src …)) 
288           [normalize in ⊢ (%→?); #H destruct (H)
289           | #x >nth_change_vec // cases (reverse ? rs0)
290             [ normalize in ⊢ (%→?); #H destruct (H)
291             | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
292         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
293           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
294           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
295             [>nth_change_vec [cases (append ???) // | @Hsrc] 
296             |@(not_to_not … Hneq) //
297             ]]
298           normalize in ⊢ (%→?); #H destruct (H) ]
299         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
300           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * *
301           >Htc >change_vec_commute // >nth_change_vec //
302           >change_vec_commute [|@sym_not_eq //] >nth_change_vec // #Hte #_ #Htb
303           #s' #rs' >Hmidta_src #H destruct (H)
304           lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
305           >change_vec_commute // >change_vec_change_vec
306           >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
307           >Hte in Htb; whd in ⊢ (%→?); #Htb >Htb %
308           [ >change_vec_change_vec >nth_change_vec //
309             >reverse_reverse <Hrs <Hmidta_src >change_vec_same <Hrs0 <Hmidta_dst
310             %
311           | >Hmidta_dst %{s'} % [%] #_
312             >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
313           ]
314         ]
315       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
316         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
317         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
318         #Htd destruct (Htd) * #te * * #_ #Hte whd in ⊢ (%→?); #Htb
319         #s1 #rs1 >Hmidta_src #H destruct (H)
320         lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
321         [ >Htb %
322         | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ]
323       ]
324     ]
325   ]
326 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
327   whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
328   lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
329   cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
330   [ #Hcurta_dst % % % // @Hcomp1 %2 //
331   | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
332     #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
333     [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
334     | >(?:tc=ta) in Htest; 
335       [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
336         #Hxx0' destruct (Hxx0') % ]
337       whd in ⊢ (??%?→?); 
338       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
339       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
340       cases (Hcomp2 … Hta_src Hta_dst) [ *
341       [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
342         [ % // | >Hcurtc % ]
343       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
344       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
345         #Hci #Hxs #Hrs0 #Htc @False_ind
346         whd in Htest:(??%?); 
347         >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest; 
348         [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
349           >nth_change_vec //]
350         >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj) 
351         [|>nth_current_chars >Htc >nth_change_vec //]
352         normalize #H destruct (H) ] ] ]
353 qed.
354
355 definition match_m ≝ λsrc,dst,sig,n.
356   whileTM … (match_step src dst sig n) 
357     (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
358
359 definition R_match_m ≝ 
360   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
361   ∀x,rs.
362   nth src ? int (niltape ?) = midtape sig [ ] x rs →
363   (current sig (nth dst (tape sig) int (niltape sig)) = None ? → 
364    right ? (nth dst (tape sig) int (niltape sig)) = [ ] → outt = int) ∧
365   (∀ls0,x0,rs0.
366    nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
367    (∃l,l1.x0::rs0 = l@x::rs@l1 ∧
368     outt = change_vec ?? 
369            (change_vec ?? int 
370              (mk_tape sig (reverse ? rs@[x]) (None ?) [ ]) src)
371            (mk_tape sig ((reverse ? (l@x::rs))@ls0) (option_hd ? l1) (tail ? l1)) dst) ∨
372     ∀l,l1.x0::rs0 ≠ l@x::rs@l1).
373
374 lemma not_sub_list_merge : 
375   ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
376 #T #a #b #H1 #H2 #l elim l normalize //
377 qed.
378
379 lemma not_sub_list_merge_2 : 
380   ∀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.
381 #T #a #b #t #H1 #H2 #l elim l //
382 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
383 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
384 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
385 qed.
386
387
388 lemma wsem_match_m : ∀src,dst,sig,n.
389 src ≠ dst → src < S n → dst < S n → 
390   match_m src dst sig n ⊫ R_match_m src dst sig n.
391 #src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
392 lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
393 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
394 [ #Hfalse #x #xs #Hmid_src
395   cases (Hfalse … Hmid_src) -Hfalse 
396   [(* current dest = None *) *
397     [ * #Hcur_dst #Houtc %
398       [#_ >Houtc //
399       | #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst; 
400         normalize in ⊢ (%→?); #H destruct (H)
401       ]
402     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
403       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
404       | #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
405         >Hrs0 >HNone cases xs0
406         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
407           @eq_f3 //
408           [ >reverse_append %
409           | >reverse_append >reverse_cons >reverse_append
410             >associative_append >associative_append % ]
411         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
412           >length_append whd in ⊢ (??%(??%)→?); >length_append
413           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
414           #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
415           >associative_plus >associative_plus 
416           #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
417           #e2 destruct (e2)
418         ]
419       ]
420     ]
421   |* #ls0 * #rs0 * #Hmid_dst #Houtc %
422     [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
423     |#ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
424      %1 %{[ ]} %{rs0} % [%] 
425      >reverse_cons >associative_append >Houtc %
426     ]
427   ]
428 |-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
429  #x #xs #Hmidta_src
430  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
431  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
432   [#Hcurta_dst % 
433     [#Hcurta_dst #Hrightta_dst whd in Htrue; >Hmidta_src in Htrue; #Htrue
434      cases (Htrue ?? (refl ??)) -Htrue #Htc
435      cut (tc = ta)
436      [ >Htc whd in match (tape_move_mono ???); whd in match (tape_write ???);
437        <(change_vec_same … ta dst (niltape ?)) in ⊢ (???%);
438        lapply Hrightta_dst lapply Hcurta_dst -Hrightta_dst -Hcurta_dst 
439        cases (nth dst ? ta (niltape ?))
440        [ #_ #_ %
441        | #r0 #rs0 #_ normalize in ⊢ (%→?); #H destruct (H)
442        | #l0 #ls0 #_ #_ %
443        | #ls #x0 #rs normalize in ⊢ (%→?); #H destruct (H) ] ]
444      -Htc #Htc destruct (Htc) #_
445      cases (IH … Hmidta_src) #Houtc #_ @Houtc //
446     |#ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst;
447      normalize in ⊢ (%→?); #H destruct (H)
448     ]
449   | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
450     #ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst; normalize in ⊢ (%→?);
451     #H destruct (H) whd in Htrue; >Hmidta_src in Htrue; #Htrue
452     cases (Htrue ?? (refl …)) -Htrue >Hmidta_dst #Htc
453     cases (true_or_false (x==c)) #eqx
454     [ lapply (\P eqx) -eqx #eqx destruct (eqx) * #s0 * whd in ⊢ (??%?→?); #Hs0
455       destruct (Hs0) #Htrue cases (Htrue (refl ??)) -Htrue
456       #xs0 * #ci * #rs' * #ls1 * #cj * #rs1 * * #Hxs #H destruct (H) #Hcicj
457       >Htc in IH; whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
458       #IH cases (IH … Hmidta_src) -IH #_ >nth_change_vec //
459       cut (∃x1,xs1.xs0@cj::rs1 = x1::xs1)
460       [ cases xs0 [ %{cj} %{rs1} % | #x1 #xs1 %{x1} %{(xs1@cj::rs1)} % ] ] * #x1 * #xs1
461       #Hxs1 >Hxs1 #IH cases (IH … (refl ??)) -IH
462       [ * #l * #l1 * #Hxs1'
463         >change_vec_commute // >change_vec_change_vec
464         #Houtc % %{(s0::l)} %{l1} % 
465         [ normalize <Hxs1' %
466         | >reverse_cons >associative_append >change_vec_commute // @Houtc ]
467       | #H %2 #l #l1 >(?:l@s0::xs@l1 = l@(s0::xs)@l1) [|%]
468         @not_sub_list_merge
469         [ #l2 >Hxs <Hxs1 % normalize #H1 lapply (cons_injective_r ????? H1)
470           >associative_append #H2 lapply (append_l2_injective ????? (refl ??) H2)
471           #H3 lapply (cons_injective_l ????? H3) #H3 >H3 in Hcicj; * /2/
472         |#t #l2 #l3 % normalize #H1 lapply (cons_injective_r ????? H1)
473          -H1 #H1 cases (H l2 l3) #H2 @H2 @H1
474         ]
475       ]
476     | #_ cases (IH x xs ?) -IH
477       [| >Htc >nth_change_vec_neq [|@sym_not_eq //] @Hmidta_src ]
478       >Htc >nth_change_vec // cases rs0
479       [ #_ #_ %2 #l #l1 cases l
480        [ normalize cases xs
481          [ cases l1
482            [ normalize % #H destruct (H) cases (\Pf eqx) /2/
483            | #tmp1 #l2 normalize % #H destruct (H) ]
484          | #tmp1 #l2 normalize % #H destruct (H) ]
485        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
486          [ normalize #H1 destruct (H1)
487          | #tmp2 #l3 normalize #H1 destruct (H1) ] ]
488       | #r1 #rs1 #_ #IH cases (IH … (refl ??)) -IH
489         [ * #l * #l1 * #Hll1 #Houtc % %{(c::l)} %{l1} % [ >Hll1 % ]
490           >Houtc >change_vec_commute // >change_vec_change_vec 
491           >change_vec_commute [|@sym_not_eq //]
492           >reverse_cons >associative_append %
493         | #Hll1 %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@Hll1]
494          #l1 % #H destruct (H) cases (\Pf eqx) /2/
495         ]
496       ]
497     ]
498   ]
499 ]
500 qed.
501
502 definition R_match_step_true_naive ≝ 
503   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
504   |left ? (nth src ? outt (niltape ?))| +
505   |option_cons ? (current ? (nth dst ? outt (niltape ?))) (right ? (nth dst ? outt (niltape ?)))| <
506   |left ? (nth src ? int (niltape ?))| +
507   |option_cons ? (current ? (nth dst ? int (niltape ?))) (right ? (nth dst ? int (niltape ?)))|.
508
509 lemma sem_match_step_termination :
510   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
511   match_step src dst sig n ⊨ 
512     [ inr ?? (inr ?? (inl … (inr ?? start_nop))) : 
513       R_match_step_true_naive src dst sig n, 
514       R_match_step_false src dst sig n ].
515 #src #dst #sig #n #Hneq #Hsrc #Hdst 
516 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
517     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
518       (sem_seq … 
519         (sem_rewind_strong ???? Hneq Hsrc Hdst) 
520         (sem_move_multi … R ?))
521       (sem_nop …))) [/2/]
522 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
523   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
524   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
525     whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
526     >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
527   | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
528     cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
529     [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
530       whd in ⊢ (%→?); * whd in ⊢ (??%?→?); 
531       >nth_current_chars >nth_current_chars >Hs >Hcurta_dst 
532       normalize in ⊢ (%→?); #H destruct (H)
533     | #s0 #Hs0
534       cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
535       cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
536       cases (true_or_false (s == s0)) #Hss0
537       [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0) 
538         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
539         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
540           >nth_current_chars >nth_current_chars >Hcurtc_dst 
541           cases (current ? (nth src …)) 
542           [normalize in ⊢ (%→?); #H destruct (H)
543           | #x >nth_change_vec [|@Hdst] cases (reverse ? rs0)
544             [ normalize in ⊢ (%→?); #H destruct (H)
545             | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
546         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
547           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
548           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
549             [>nth_change_vec [cases (append ???) // | @Hsrc] 
550             |@(not_to_not … Hneq) //
551             ]]
552           normalize in ⊢ (%→?); #H destruct (H) ]
553         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
554           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * * *
555           >Htc >change_vec_commute [|//] >nth_change_vec [|//]
556           >change_vec_commute [|@sym_not_eq //] >nth_change_vec [|//]
557           cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb
558           destruct (Hlsalsb)  *
559           [ #Hte #_ #_ <(reverse_reverse … ls) in Hte; <(reverse_reverse … lsa)
560             cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen' 
561             @(list_cases2 … Hlen')
562             [ #H1 #H2 >H1 >H2 -H1 -H2 normalize in match (reverse ? [ ]); #Hte #_
563               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
564               >change_vec_commute [|//] >change_vec_change_vec
565               >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
566               >Hte whd in ⊢ (%→?); >change_vec_change_vec >nth_change_vec [|//] 
567               >reverse_reverse #Htb
568               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)
569               [ >Htb @eq_f3 // cases (xs@cj::rs0') // ]
570               -Htb #Htb >Htb whd >nth_change_vec [|//]
571               >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec [|//]
572               >right_mk_tape [|cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H)]
573               normalize in match (left ??);
574               >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand
575               whd in match (option_cons ???); >Hrs0
576               normalize in ⊢ (?(?%)%); //
577             | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
578               >reverse_cons >reverse_cons #Hte
579               lapply (Hte ci hdb (reverse ? xs@s0::reverse ? tlb) rs'' ?
580                        lsb cj hda (reverse ? xs@s0::reverse ? tla) rs0' ??)
581               [ /2 by cons_injective_l, nil/
582               | >length_append >length_append @eq_f @(eq_f ?? S)
583                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
584                 >length_reverse >length_reverse destruct (Hlen') //
585               | /2 by refl, trans_eq/ ] -Hte
586               #Hte #_  whd in ⊢ (%→?); #Htb
587               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
588                         (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) 
589                         (midtape ? [ ] hdb (reverse sig (reverse sig xs@s0::reverse sig tlb)@ci::rs'')) src)
590               [ >Htb >Hte >nth_change_vec // >change_vec_change_vec >change_vec_commute [|//]
591                 >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
592                 >change_vec_change_vec >change_vec_commute [|//]
593                 @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0') // ]
594               -Htb #Htb >Htb whd 
595               >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
596               >right_mk_tape 
597               [| cases (reverse sig (reverse sig xs@s0::reverse sig tla))
598                  [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
599               >Hmidta_src >Hmidta_dst 
600               whd in match (left ??); whd in match (left ??); whd in match (right ??);
601               >current_mk_tape <opt_cons_tail_expand whd in match (option_cons ???);
602               >Hrs0 >length_append whd in ⊢ (??(??%)); >length_append >length_reverse
603               >length_append >commutative_plus in match (|reverse ??| + ?);
604               whd in match (|?::?|); >length_reverse >length_reverse
605               <(length_reverse ? ls) <Hlen' >H1 normalize // ]
606          | #_ #Hte #_ <(reverse_reverse … ls0) in Hte; <(reverse_reverse … lsa)
607             cut (|reverse ? lsa| = |reverse ? ls0|) [ // ] #Hlen' 
608             @(list_cases2 … Hlen')
609             [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #Hte
610               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
611               >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] 
612               >change_vec_change_vec #Hte #_
613               >Hte whd in ⊢ (%→?); >nth_change_vec [|//] >reverse_reverse #Htb 
614               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (mk_tape ? [s0] (option_hd ? (xs@cj::rs0')) (tail ? (xs@cj::rs0'))) dst)
615                            (midtape ? lsb s0 (xs@ci::rs'')) src)
616               [ >Htb >change_vec_change_vec >change_vec_commute [|//]
617                 @eq_f3 // <Hrs0 cases rs0 // ]
618               -Htb #Htb >Htb whd >nth_change_vec [|//]
619               >nth_change_vec_neq [|//] >nth_change_vec [|//]
620               >right_mk_tape 
621               [| cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] 
622               normalize in match (left ??);
623               >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand >Hrs0
624               >length_append normalize >length_append >length_append
625               <(reverse_reverse ? lsa) >H1 normalize //
626             | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
627               >reverse_cons >reverse_cons #Hte
628               lapply (Hte cj hdb (reverse ? xs@s0::reverse ? tlb) rs0' ?
629                        lsb ci hda (reverse ? xs@s0::reverse ? tla) rs'' ??)
630               [ /2 by cons_injective_l, nil/
631               | >length_append >length_append @eq_f @(eq_f ?? S)
632                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
633                 >length_reverse >length_reverse destruct (Hlen') //
634               | /2 by refl, trans_eq/ ] -Hte
635               #Hte #_ whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb
636               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
637                         (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) 
638                         (midtape ? lsb hda (reverse sig (reverse sig xs@s0::reverse sig tla)@ci::rs'')) src)
639               [ >Htb >change_vec_change_vec >change_vec_commute [|//]
640                 >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
641                 >change_vec_change_vec >change_vec_commute [|//]
642                 @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0') // ]
643               -Htb #Htb >Htb whd 
644               >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
645               >right_mk_tape
646               [| cases (reverse sig (reverse sig xs@s0::reverse sig tlb))
647                  [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
648               >Hmidta_src >Hmidta_dst 
649               whd in match (left ??); whd in match (left ??); whd in match (right ??);
650               >current_mk_tape <opt_cons_tail_expand
651               whd in match (option_cons ???);
652               >Hrs0 >length_append whd in ⊢ (??(??%)); >length_append >length_reverse
653               >length_append >commutative_plus in match (|reverse ??| + ?);
654               whd in match (|?::?|); >length_reverse >length_reverse
655               <(length_reverse ? lsa) >Hlen' >H2 >length_append
656               normalize //
657             ]
658           ]
659         ]
660       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
661         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
662         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
663         #Htd destruct (Htd) * #te * * * * >Hmidta_src >Hmidta_dst 
664         cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb
665         destruct (Hlsalsb)
666         [ <(reverse_reverse … ls) <(reverse_reverse … lsa)
667           cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen'
668           @(list_cases2 … Hlen')
669           [ #H1 #H2 >H1 >H2 -H1 -H2 #_ #_ normalize in match (reverse ? [ ]); #Hte #_
670             lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte)
671             whd in ⊢ (%→?); >Hmidta_dst #Htb
672             cut (tb = change_vec ?? ta (mk_tape ? (s0::lsa@lsb) (option_hd ? rs0) (tail ? rs0)) dst)
673             [ >Htb cases rs0 // ]
674             -Htb #Htb >Htb whd >nth_change_vec [|//]
675             >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst
676             >right_mk_tape 
677             [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H)] ]
678             normalize in match (left ??); normalize in match (right ??);
679             >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand
680             normalize //
681           | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
682             >reverse_cons >reverse_cons >associative_append #Hte
683             lapply (Hte ???? (refl ??) ? s0 ? (reverse ? tla) ?? (refl ??))
684             [ >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) //
685               normalize #H destruct (H) // ] #Hte #_ #_ #_
686             whd in ⊢ (%→?); >Hte >change_vec_change_vec >nth_change_vec // #Htb
687             cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
688                       (mk_tape sig (hda::lsb) (option_hd ? (reverse sig (reverse sig tla)@s0::rs0)) (tail ? (reverse sig (reverse sig tla)@s0::rs0))) dst) 
689                       (midtape ? [ ] hdb (reverse sig (reverse sig tlb)@s::rs)) src)
690             [ >Htb >change_vec_commute [|//] @eq_f3 // cases (reverse sig (reverse sig tla)@s0::rs0) // ]
691             -Htb #Htb >Htb whd 
692             >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
693             >right_mk_tape 
694             [| cases (reverse sig (reverse sig tla))
695                [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
696             >Hmidta_src >Hmidta_dst 
697             whd in match (left ??); whd in match (left ??); whd in match (right ??);
698             >current_mk_tape <opt_cons_tail_expand >length_append
699             >length_reverse >length_reverse <(length_reverse ? ls) <Hlen'
700             >H1 normalize // ]
701        | #_ <(reverse_reverse … ls0) <(reverse_reverse … lsa)
702          cut (|reverse ? lsa| = |reverse ? ls0|) [ // ] #Hlen' 
703          @(list_cases2 … Hlen')
704          [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #_ #_ #Hte
705            lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte)
706            whd in ⊢ (%→?); #Htb whd >Hmidta_dst
707            cut (tb = change_vec (tape sig) (S n) ta (mk_tape ? (s0::ls0) (option_hd ? rs0) (tail ? rs0)) dst)
708            [ >Htb >Hmidta_dst cases rs0 // ]
709            -Htb #Htb >Htb whd >nth_change_vec [|//]
710            >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst
711            >current_mk_tape >right_mk_tape 
712            [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H) ]]
713            normalize in ⊢ (??%); <opt_cons_tail_expand
714            normalize //
715          | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
716            >reverse_cons >reverse_cons #Hte #_ #_
717            lapply (Hte s0 hdb (reverse ? tlb) rs0 ?
718                     lsb s hda (reverse ? tla) rs ??)
719            [ /2 by cons_injective_l, nil/
720            | >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) //
721              normalize #H destruct (H) //
722            | /2 by refl, trans_eq/ ] -Hte
723            #Hte whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb
724            cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
725                      (mk_tape sig [hdb] (option_hd ? (reverse sig (reverse sig tlb)@s0::rs0)) (tail ? (reverse sig (reverse sig tlb)@s0::rs0))) dst) 
726                      (midtape ? lsb hda (reverse sig (reverse sig tla)@s::rs)) src)
727            [ >Htb >change_vec_commute [|//] >change_vec_change_vec
728              @eq_f3 // cases (reverse sig (reverse sig tlb)@s0::rs0) // ]
729            -Htb #Htb >Htb whd 
730            >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
731            >right_mk_tape
732            [| cases (reverse ? (reverse ? tlb)) [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
733            >Hmidta_src >Hmidta_dst 
734            whd in match (left ??); whd in match (left ??); whd in match (right ??);
735            >current_mk_tape <opt_cons_tail_expand >length_append
736            normalize in ⊢ (??%); >length_append >reverse_reverse
737            <(length_reverse ? lsa) >Hlen' >H2 normalize //
738          ]
739        ]
740      ]
741    ]
742  ]
743 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
744   whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
745   lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
746   cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
747   [ #Hcurta_dst % % % // @Hcomp1 %2 //
748   | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
749     #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
750     [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
751     | >(?:tc=ta) in Htest; 
752       [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
753         #Hxx0' destruct (Hxx0') % ]
754       whd in ⊢ (??%?→?); 
755       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
756       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
757       cases (Hcomp2 … Hta_src Hta_dst) [ *
758       [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
759         [ % // | >Hcurtc % ]
760       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
761       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
762         #Hci #Hxs #Hrs0 #Htc @False_ind
763         whd in Htest:(??%?); 
764         >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest; 
765         [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
766           >nth_change_vec //]
767         >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj) 
768         [|>nth_current_chars >Htc >nth_change_vec //]
769         normalize #H destruct (H) ] ] ]
770 qed.
771
772 (* lemma WF_to_WF_f : ∀A,B,R,f,b. WF A R (f b) → WF B (λx,y.R (f x) (f y)) b. *)
773 let rec WF_to_WF_f A B R f b (Hwf: WF A R (f b)) on Hwf: WF B (λx,y.R (f x) (f y)) b ≝ 
774   match Hwf return (λa0,r.f b = a0 → WF B (λx,y:B. R (f x) (f y)) b) with
775   [ wf a Hwfa ⇒ λHeq.? ] (refl ??).
776 % #b1 #HRb @WF_to_WF_f @Hwfa <Heq @HRb
777 qed.
778
779 lemma lt_WF : ∀n.WF ? lt n.
780 #n @(nat_elim1 n) -n #n #IH % @IH
781 qed.
782
783 lemma terminate_match_m :
784   ∀src,dst,sig,n,t.
785   src ≠ dst → src < S n → dst < S n → 
786   match_m src dst sig n ↓ t.
787 #src #dst #sig #n #ta #Hneq #Hsrc #Hdst
788 @(terminate_while … (sem_match_step_termination src dst sig n Hneq Hsrc Hdst)) //
789 letin f ≝ (λt0:Vector (tape sig) (S n).|left ? (nth src (tape ?) t0 (niltape ?))|
790     +|option_cons ? (current ? (nth dst (tape ?) t0 (niltape ?)))
791       (right ? (nth dst (tape ?) t0 (niltape ?)))|)
792 change with (λx,y.f x < f y) in ⊢ (??%?); @WF_to_WF_f @lt_WF
793 qed.
794
795 lemma sem_match_m : ∀src,dst,sig,n.
796 src ≠ dst → src < S n → dst < S n → 
797   match_m src dst sig n \vDash R_match_m src dst sig n.
798 #src #dst #sig #n #Hneq #Hsrc #Hdst @WRealize_to_Realize [/2/| @wsem_match_m // ]
799 qed.