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