1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "turing/multi_universal/compare.ma".
16 include "turing/multi_universal/par_test.ma".
17 include "turing/multi_universal/moves_2.ma".
19 definition Rtc_multi_true ≝
20 λalpha,test,n,i.λt1,t2:Vector ? (S n).
21 (∃c. current alpha (nth i ? t1 (niltape ?)) = Some ? c ∧ test c = true) ∧ t2 = t1.
23 definition Rtc_multi_false ≝
24 λalpha,test,n,i.λt1,t2:Vector ? (S n).
25 (∀c. current alpha (nth i ? t1 (niltape ?)) = Some ? c → test c = false) ∧ t2 = t1.
27 lemma sem_test_char_multi :
28 ∀alpha,test,n,i.i ≤ n →
29 inject_TM ? (test_char ? test) n i ⊨
30 [ tc_true : Rtc_multi_true alpha test n i, Rtc_multi_false alpha test n i ].
31 #alpha #test #n #i #Hin #int
32 cases (acc_sem_inject … Hin (sem_test_char alpha test) int)
33 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
35 | #Hqtrue lapply (Htrue Hqtrue) * * * #c *
36 #Hcur #Htestc #Hnth_i #Hnth_j %
38 | @(eq_vec … (niltape ?)) #i0 #Hi0
39 cases (decidable_eq_nat i0 i) #Hi0i
41 | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
42 | #Hqfalse lapply (Hfalse Hqfalse) * * #Htestc #Hnth_i #Hnth_j %
44 | @(eq_vec … (niltape ?)) #i0 #Hi0
45 cases (decidable_eq_nat i0 i) #Hi0i
47 | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
50 definition Rm_test_null_true ≝
51 λalpha,n,i.λt1,t2:Vector ? (S n).
52 current alpha (nth i ? t1 (niltape ?)) ≠ None ? ∧ t2 = t1.
54 definition Rm_test_null_false ≝
55 λalpha,n,i.λt1,t2:Vector ? (S n).
56 current alpha (nth i ? t1 (niltape ?)) = None ? ∧ t2 = t1.
58 lemma sem_test_null_multi : ∀alpha,n,i.i ≤ n →
59 inject_TM ? (test_null ?) n i ⊨
60 [ tc_true : Rm_test_null_true alpha n i, Rm_test_null_false alpha n i ].
61 #alpha #n #i #Hin #int
62 cases (acc_sem_inject … Hin (sem_test_null alpha) int)
63 #k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
65 | #Hqtrue lapply (Htrue Hqtrue) * * #Hcur #Hnth_i #Hnth_j % //
66 @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) #Hi0i
67 [ >Hi0i @sym_eq @Hnth_i | @sym_eq @Hnth_j @sym_not_eq // ] ]
68 | #Hqfalse lapply (Hfalse Hqfalse) * * #Hcur #Hnth_i #Hnth_j %
70 | @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) //
71 #Hi0i @sym_eq @Hnth_j @sym_not_eq // ] ]
74 definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
75 match (nth src (option sig) v (None ?)) with
77 | Some x ⇒ notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
79 definition mmove_states ≝ initN 2.
81 definition mmove0 : mmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
82 definition mmove1 : mmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
84 definition trans_mmove ≝
86 λp:mmove_states × (Vector (option sig) (S n)).
87 let 〈q,a〉 ≝ p in match (pi1 … q) with
88 [ O ⇒ 〈mmove1,change_vec ? (S n) (null_action ? n) (〈None ?,D〉) i〉
89 | S _ ⇒ 〈mmove1,null_action sig n〉 ].
93 mk_mTM sig n mmove_states (trans_mmove i sig n D)
94 mmove0 (λq.q == mmove1).
97 λalpha,n,i,D.λt1,t2:Vector ? (S n).
98 t2 = change_vec ? (S n) t1 (tape_move alpha (nth i ? t1 (niltape ?)) D) i.
100 lemma sem_move_multi :
102 mmove i alpha n D ⊨ Rm_multi alpha n i D.
103 #alpha #n #i #D #Hin #int %{2}
104 %{(mk_mconfig ? mmove_states n mmove1 ?)}
106 [ whd in ⊢ (??%?); @eq_f whd in ⊢ (??%?); @eq_f %
107 | whd >tape_move_multi_def
108 <(change_vec_same … (ctapes …) i (niltape ?))
109 >pmap_change <tape_move_multi_def >tape_move_null_action % ] ]
112 definition rewind ≝ λsrc,dst,sig,n.
113 parmove src dst sig n L · mmove src sig n R · mmove dst sig n R.
115 definition R_rewind ≝ λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
117 nth src ? int (niltape ?) = midtape sig (xs@[x0]) x rs →
118 ∀ls0,y,y0,target,rs0.|xs| = |target| →
119 nth dst ? int (niltape ?) = midtape sig (target@y0::ls0) y rs0 →
121 (change_vec ?? int (midtape sig [] x0 (reverse ? xs@x::rs)) src)
122 (midtape sig ls0 y0 (reverse ? target@y::rs0)) dst) ∧
123 (∀x,rs.nth src ? int (niltape ?) = midtape sig [] x rs →
124 ∀ls0,y,rs0.nth dst ? int (niltape ?) = midtape sig ls0 y rs0 →
128 theorem accRealize_to_Realize :
129 ∀sig,n.∀M:mTM sig n.∀Rtrue,Rfalse,acc.
130 M ⊨ [ acc: Rtrue, Rfalse ] → M ⊨ Rtrue ∪ Rfalse.
131 #sig #n #M #Rtrue #Rfalse #acc #HR #t
132 cases (HR t) #k * #outc * * #Hloop
133 #Htrue #Hfalse %{k} %{outc} % //
134 cases (true_or_false (cstate sig (states sig n M) n outc == acc)) #Hcase
135 [ % @Htrue @(\P Hcase) | %2 @Hfalse @(\Pf Hcase) ]
139 lemma sem_rewind : ∀src,dst,sig,n.
140 src ≠ dst → src < S n → dst < S n →
141 rewind src dst sig n ⊨ R_rewind src dst sig n.
142 #src #dst #sig #n #Hneq #Hsrc #Hdst
143 @(sem_seq_app sig n ????? (sem_parmoveL src dst sig n Hneq Hsrc Hdst) ?)
144 [| @(sem_seq_app sig n ????? (sem_move_multi … R ?) (sem_move_multi … R ?)) //
146 #ta #tb * #tc * * #Htc #_ * #td * whd in ⊢ (%→%→?); #Htd #Htb %
147 [ #x #x0 #xs #rs #Hmidta_src #ls0 #y #y0 #target #rs0 #Hlen #Hmidta_dst
148 >(Htc ??? Hmidta_src ls0 y (target@[y0]) rs0 ??) in Htd;
150 |>length_append >length_append >Hlen % ]
151 >change_vec_commute [|@sym_not_eq //]
152 >change_vec_change_vec
153 >nth_change_vec_neq [|@sym_not_eq //]
154 >nth_change_vec // >reverse_append >reverse_single
155 >reverse_append >reverse_single normalize in match (tape_move ???);
156 >rev_append_def >append_nil #Htd >Htd in Htb;
157 >change_vec_change_vec >nth_change_vec //
158 cases ls0 [|#l1 #ls1] normalize in match (tape_move ???); //
159 | #x #rs #Hmidta_src #ls0 #y #rs0 #Hmidta_dst
160 lapply (Htc … Hmidta_src … (refl ??) Hmidta_dst) -Htc #Htc >Htc in Htd;
161 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
162 >nth_change_vec_neq [|@sym_not_eq //]
163 >nth_change_vec // lapply (refl ? ls0) cases ls0 in ⊢ (???%→%);
164 [ #Hls0 #Htd >Htd in Htb;
165 >nth_change_vec // >change_vec_change_vec
166 whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
167 <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
168 | #l1 #ls1 #Hls0 #Htd >Htd in Htb;
169 >nth_change_vec // >change_vec_change_vec
170 whd in match (tape_move ???);whd in match (tape_move ???); <Hmidta_src
171 <Hls0 <Hmidta_dst >change_vec_same >change_vec_same //
175 definition match_step ≝ λsrc,dst,sig,n.
176 compare src dst sig n ·
177 (ifTM ?? (partest sig n (match_test src dst sig ?))
179 (rewind src dst sig n · (inject_TM ? (move_r ?) n dst)))
183 (* we assume the src is a midtape
185 if the dst is out of bounds (outt = int)
186 or dst.right is shorter than src.right (outt.current → None)
187 or src.right is a prefix of dst.right (out = just right of the common prefix) *)
188 definition R_match_step_false ≝
189 λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
191 nth src ? int (niltape ?) = midtape sig ls x xs →
192 ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
193 (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
196 (change_vec ?? int (mk_tape sig (reverse ? rs0@x::ls) (option_hd ? xs0) (tail ? xs0)) src)
197 (mk_tape ? (reverse ? rs0@x::ls0) (None ?) [ ]) dst) ∨
199 nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
203 (change_vec ?? int (mk_tape sig (reverse ? xs@x::ls) (None ?) [ ]) src)
204 (mk_tape sig (reverse ? xs@x::ls0) (option_hd ? rs0) (tail ? rs0)) dst).
207 we assume the src is a midtape [ ] s rs
209 then dst.current = Some ? s1
210 and if s ≠ s1 then outt = int.dst.move_right()
212 then int.src.right and int.dst.right have a common prefix
213 and the heads of their suffixes are different
214 and outt = int.dst.move_right().
217 definition R_match_step_true ≝
218 λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
219 ∀s,rs.nth src ? int (niltape ?) = midtape ? [ ] s rs →
220 outt = change_vec ?? int
221 (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst ∧
222 (∃s0.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s0 ∧
224 ∃xs,ci,rs',ls0,cj,rs0.
226 nth dst ? int (niltape ?) = midtape sig ls0 s (xs@cj::rs0) ∧
229 lemma sem_match_step :
230 ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n →
231 match_step src dst sig n ⊨
232 [ inr ?? (inr ?? (inl … (inr ?? start_nop))) :
233 R_match_step_true src dst sig n,
234 R_match_step_false src dst sig n ].
235 #src #dst #sig #n #Hneq #Hsrc #Hdst
236 @(acc_sem_seq_app sig n … (sem_compare src dst sig n Hneq Hsrc Hdst)
237 (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
239 (sem_rewind ???? Hneq Hsrc Hdst)
240 (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
242 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
243 cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
244 [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
245 whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
246 >nth_current_chars >Hcurta_src normalize in ⊢ (%→?); #H destruct (H)
247 | #s #Hs lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
248 cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→%);
249 [ #Hcurta_dst #Hcomp #_ * #td * >Hcomp [| %2 %]
250 whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
251 >nth_current_chars >nth_current_chars >Hs >Hcurta_dst
252 normalize in ⊢ (%→?); #H destruct (H)
254 cases (current_to_midtape … Hs) #ls * #rs #Hmidta_src >Hmidta_src
255 cases (current_to_midtape … Hs0) #ls0 * #rs0 #Hmidta_dst >Hmidta_dst
256 cases (true_or_false (s == s0)) #Hss0
257 [ lapply (\P Hss0) -Hss0 #Hss0 destruct (Hss0)
258 #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
259 [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
260 >nth_current_chars >nth_current_chars >Hcurtc_dst
261 cases (current ? (nth src …))
262 [normalize in ⊢ (%→?); #H destruct (H)
263 | #x >nth_change_vec // cases (reverse ? rs0)
264 [ normalize in ⊢ (%→?); #H destruct (H)
265 | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
266 | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
267 >(?:nth src ? (current_chars ?? tc) (None ?) = None ?)
268 [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq
269 [>nth_change_vec [cases (append ???) // | @Hsrc]
270 |@(not_to_not … Hneq) //
272 normalize in ⊢ (%→?); #H destruct (H) ]
273 | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
274 #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * *
275 >Htc >change_vec_commute // >nth_change_vec //
276 >change_vec_commute [|@sym_not_eq //] >nth_change_vec // #Hte #_ #Htb
277 #s' #rs' >Hmidta_src #H destruct (H)
278 lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
279 >change_vec_commute // >change_vec_change_vec
280 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
281 >Hte in Htb; * * #_ >nth_change_vec // #Htb1
282 lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 %
283 [ @(eq_vec … (niltape ?)) #i #Hi
284 cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
285 [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst
286 >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] %
287 | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)]
288 >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%);
289 <Hrs <Hmidta_src [|@(\Pf Hdsti)] >change_vec_same % ]
290 | >Hmidta_dst %{s'} % [%] #_
291 >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % //
294 | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta)
295 [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
296 -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_
297 #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2
298 #s1 #rs1 >Hmidta_src #H destruct (H)
299 lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
300 [ @(eq_vec … (niltape ?)) #i #Hi
301 cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
302 [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
303 cases rs0 [|#r2 #rs2] %
304 | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
305 | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ]
309 | #ta #tb #tc * #Hcomp1 #Hcomp2 * #td * * #Htest #Htd destruct (Htd)
310 whd in ⊢ (%→?); #Htb destruct (Htb) #ls #x #xs #Hta_src
311 lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
312 cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
313 [ #Hcurta_dst % % % // @Hcomp1 %2 //
314 | #x0 #Hcurta_dst cases (current_to_midtape … Hcurta_dst) -Hcurta_dst
315 #ls0 * #rs0 #Hta_dst cases (true_or_false (x == x0)) #Hxx0
316 [ lapply (\P Hxx0) -Hxx0 #Hxx0 destruct (Hxx0)
317 | >(?:tc=ta) in Htest;
318 [|@Hcomp1 % % >Hta_src >Hta_dst @(not_to_not ??? (\Pf Hxx0)) normalize
319 #Hxx0' destruct (Hxx0') % ]
321 >nth_current_chars >Hta_src >nth_current_chars >Hta_dst
322 whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
323 cases (Hcomp2 … Hta_src Hta_dst) [ *
324 [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
326 | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
327 | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
328 #Hci #Hxs #Hrs0 #Htc @False_ind
330 >(?:nth src ? (current_chars ?? tc) (None ?) = Some ? ci) in Htest;
331 [|>nth_current_chars >Htc >nth_change_vec_neq [|@(not_to_not … Hneq) //]
333 >(?:nth dst ? (current_chars ?? tc) (None ?) = Some ? cj)
334 [|>nth_current_chars >Htc >nth_change_vec //]
335 normalize #H destruct (H) ] ] ]
338 definition match_m ≝ λsrc,dst,sig,n.
339 whileTM … (match_step src dst sig n)
340 (inr ?? (inr ?? (inl … (inr ?? start_nop)))).
342 definition R_match_m ≝
343 λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
345 nth src ? int (niltape ?) = midtape sig [ ] x rs →
346 (current sig (nth dst (tape sig) int (niltape sig)) = None ? →
347 right ? (nth dst (tape sig) int (niltape sig)) = [ ] → outt = int) ∧
349 nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
350 (∃l,l1.x0::rs0 = l@x::rs@l1 ∧
353 (mk_tape sig (reverse ? rs@[x]) (None ?) [ ]) src)
354 (mk_tape sig ((reverse ? (l@x::rs))@ls0) (option_hd ? l1) (tail ? l1)) dst) ∨
355 ∀l,l1.x0::rs0 ≠ l@x::rs@l1).
357 lemma not_sub_list_merge :
358 ∀T.∀a,b:list T. (∀l1.a ≠ b@l1) → (∀t,l,l1.a ≠ t::l@b@l1) → ∀l,l1.a ≠ l@b@l1.
359 #T #a #b #H1 #H2 #l elim l normalize //
362 lemma not_sub_list_merge_2 :
363 ∀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.
364 #T #a #b #t #H1 #H2 #l elim l //
365 #t0 #l1 #IH #l2 cases (true_or_false (t == t0)) #Htt0
366 [ >(\P Htt0) % normalize #H destruct (H) cases (H2 l1 l2) /2/
367 | normalize % #H destruct (H) cases (\Pf Htt0) /2/ ]
371 lemma wsem_match_m : ∀src,dst,sig,n.
372 src ≠ dst → src < S n → dst < S n →
373 match_m src dst sig n ⊫ R_match_m src dst sig n.
374 #src #dst #sig #n #Hneq #Hsrc #Hdst #ta #k #outc #Hloop
375 lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) //
376 -Hloop * #tb * #Hstar @(star_ind_l ??????? Hstar) -Hstar
377 [ #Hfalse #x #xs #Hmid_src
378 cases (Hfalse … Hmid_src) -Hfalse
379 [(* current dest = None *) *
380 [ * #Hcur_dst #Houtc %
382 | #ls0 #x0 #rs0 #Hmid_dst >Hmid_dst in Hcur_dst;
383 normalize in ⊢ (%→?); #H destruct (H)
385 | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
386 [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
387 | #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
388 >Hrs0 >HNone cases xs0
389 [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
392 | >reverse_append >reverse_cons >reverse_append
393 >associative_append >associative_append % ]
394 | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
395 >length_append whd in ⊢ (??%(??%)→?); >length_append
396 >length_append normalize >commutative_plus whd in ⊢ (???%→?);
397 #H destruct (H) lapply e0 >(plus_n_O (|rs1|)) in ⊢ (??%?→?);
398 >associative_plus >associative_plus
399 #e1 lapply (injective_plus_r ??? e1) whd in ⊢ (???%→?);
404 |* #ls0 * #rs0 * #Hmid_dst #Houtc %
405 [ >Hmid_dst normalize in ⊢ (%→?); #H destruct (H)
406 |#ls1 #x1 #rs1 >Hmid_dst #H destruct (H)
407 %1 %{[ ]} %{rs0} % [%]
408 >reverse_cons >associative_append >Houtc %
411 |-ta #ta #tc #Htrue #Hstar #IH #Hout lapply (IH Hout) -IH -Hout #IH whd
413 lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
414 cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?);
416 [#Hcurta_dst #Hrightta_dst whd in Htrue; >Hmidta_src in Htrue; #Htrue
417 cases (Htrue ?? (refl ??)) -Htrue #Htc
419 [ >Htc whd in match (tape_move_mono ???); whd in match (tape_write ???);
420 <(change_vec_same … ta dst (niltape ?)) in ⊢ (???%);
421 lapply Hrightta_dst lapply Hcurta_dst -Hrightta_dst -Hcurta_dst
422 cases (nth dst ? ta (niltape ?))
424 | #r0 #rs0 #_ normalize in ⊢ (%→?); #H destruct (H)
426 | #ls #x0 #rs normalize in ⊢ (%→?); #H destruct (H) ] ]
427 -Htc #Htc destruct (Htc) #_
428 cases (IH … Hmidta_src) #Houtc #_ @Houtc //
429 |#ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst;
430 normalize in ⊢ (%→?); #H destruct (H)
432 | #c #Hcurta_dst % [ >Hcurta_dst #H destruct (H) ]
433 #ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst; normalize in ⊢ (%→?);
434 #H destruct (H) whd in Htrue; >Hmidta_src in Htrue; #Htrue
435 cases (Htrue ?? (refl …)) -Htrue >Hmidta_dst #Htc
436 cases (true_or_false (x==c)) #eqx
437 [ lapply (\P eqx) -eqx #eqx destruct (eqx) * #s0 * whd in ⊢ (??%?→?); #Hs0
438 destruct (Hs0) #Htrue cases (Htrue (refl ??)) -Htrue
439 #xs0 * #ci * #rs' * #ls1 * #cj * #rs1 * * #Hxs #H destruct (H) #Hcicj
440 >Htc in IH; whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
441 #IH cases (IH … Hmidta_src) -IH #_ >nth_change_vec //
442 cut (∃x1,xs1.xs0@cj::rs1 = x1::xs1)
443 [ cases xs0 [ %{cj} %{rs1} % | #x1 #xs1 %{x1} %{(xs1@cj::rs1)} % ] ] * #x1 * #xs1
444 #Hxs1 >Hxs1 #IH cases (IH … (refl ??)) -IH
445 [ * #l * #l1 * #Hxs1'
446 >change_vec_commute // >change_vec_change_vec
447 #Houtc % %{(s0::l)} %{l1} %
449 | >reverse_cons >associative_append >change_vec_commute // @Houtc ]
450 | #H %2 #l #l1 >(?:l@s0::xs@l1 = l@(s0::xs)@l1) [|%]
452 [ #l2 >Hxs <Hxs1 % normalize #H1 lapply (cons_injective_r ????? H1)
453 >associative_append #H2 lapply (append_l2_injective ????? (refl ??) H2)
454 #H3 lapply (cons_injective_l ????? H3) #H3 >H3 in Hcicj; * /2/
455 |#t #l2 #l3 % normalize #H1 lapply (cons_injective_r ????? H1)
456 -H1 #H1 cases (H l2 l3) #H2 @H2 @H1
459 | #_ cases (IH x xs ?) -IH
460 [| >Htc >nth_change_vec_neq [|@sym_not_eq //] @Hmidta_src ]
461 >Htc >nth_change_vec // cases rs0
462 [ #_ #_ %2 #l #l1 cases l
465 [ normalize % #H destruct (H) cases (\Pf eqx) /2/
466 | #tmp1 #l2 normalize % #H destruct (H) ]
467 | #tmp1 #l2 normalize % #H destruct (H) ]
468 | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
469 [ normalize #H1 destruct (H1)
470 | #tmp2 #l3 normalize #H1 destruct (H1) ] ]
471 | #r1 #rs1 #_ #IH cases (IH … (refl ??)) -IH
472 [ * #l * #l1 * #Hll1 #Houtc % %{(c::l)} %{l1} % [ >Hll1 % ]
473 >Houtc >change_vec_commute // >change_vec_change_vec
474 >change_vec_commute [|@sym_not_eq //]
475 >reverse_cons >associative_append %
476 | #Hll1 %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@Hll1]
477 #l1 % #H destruct (H) cases (\Pf eqx) /2/
485 definition Pre_match_m ≝
486 λsrc,sig,n.λt: Vector (tape sig) (S n).
488 nth src (tape sig) t (niltape sig) = midtape ? [] x xs.
490 lemma terminate_match_m :
492 src ≠ dst → src < S n → dst < S n →
493 Pre_match_m src sig n t →
494 match_m src dst sig n ↓ t.
495 #src #dst #sig #n #t #Hneq #Hsrc #Hdst * #start * #xs
497 @(terminate_while … (sem_match_step src dst sig n Hneq Hsrc Hdst)) //
498 <(change_vec_same … t dst (niltape ?))
499 lapply (refl ? (nth dst (tape sig) t (niltape ?)))
500 cases (nth dst (tape sig) t (niltape ?)) in ⊢ (???%→?);
501 [ #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
502 >Hmid_src #HR cases (HR ?? (refl ??)) -HR
503 >nth_change_vec // >Htape_dst #_ * #s0 * normalize in ⊢ (%→?); #H destruct (H)
504 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
505 >Hmid_src #HR cases (HR ?? (refl ??)) -HR
506 >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
507 * #s0 * #H destruct (H)
508 | #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
509 >Hmid_src #HR cases (HR ?? (refl ??)) -HR
510 >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
511 * #s0 * #H destruct (H)
512 | #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs
513 [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
514 >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ?? (refl ??)) -HR
515 >change_vec_change_vec #Ht1 #_ % #t2 whd in ⊢ (%→?);
516 >Ht1 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src #HR
517 cases (HR ?? (refl ??)) -HR #_
518 >nth_change_vec // * #s1 * normalize in ⊢ (%→?); #H destruct (H)
519 |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?);
520 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src
521 #Htrue cases (Htrue ?? (refl ??)) -Htrue >change_vec_change_vec
522 >nth_change_vec // >Hmid_dst whd in match (tape_move_mono ???); #Ht1
523 * #s0 * whd in ⊢ (??%?→?); #H destruct (H) #_ >Ht1
524 lapply (IH t1 ? (s0::ls) r0 ?)
525 [ >Ht1 >nth_change_vec //
526 | >Ht1 >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
527 | >Ht1 >nth_change_vec // ]