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