]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/match.ma
match now only uses the new move operation
[helm.git] / matita / matita / lib / turing / multi_universal / match.ma
index 74fd89cdf93b862f609a7da3cf40f5aedf7727d9..93e6b340df1d5fb0a35c72f247dff132796979be 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "turing/simple_machines.ma".
 include "turing/multi_universal/compare.ma".
 include "turing/multi_universal/par_test.ma".
 include "turing/multi_universal/moves_2.ma".
 
-definition Rtc_multi_true ≝ 
-  λalpha,test,n,i.λt1,t2:Vector ? (S n).
-   (∃c. current alpha (nth i ? t1 (niltape ?)) = Some ? c ∧ test c = true) ∧ t2 = t1.
-   
-definition Rtc_multi_false ≝ 
-  λalpha,test,n,i.λt1,t2:Vector ? (S n).
-    (∀c. current alpha (nth i ? t1 (niltape ?)) = Some ? c → test c = false) ∧ t2 = t1.
+lemma eq_vec_change_vec : ∀sig,n.∀v1,v2:Vector sig n.∀i,t,d.
+  nth i ? v2 d = t → 
+  (∀j.i ≠ j → nth j ? v1 d = nth j ? v2 d) → 
+  v2 = change_vec ?? v1 t i.
+#sig #n #v1 #v2 #i #t #d #H1 #H2 @(eq_vec … d)
+#i0 #Hlt cases (decidable_eq_nat i0 i) #Hii0
+[ >Hii0 >nth_change_vec //
+| >nth_change_vec_neq [|@sym_not_eq //] @sym_eq @H2 @sym_not_eq // ]
+qed.
 
-lemma sem_test_char_multi :
-  ∀alpha,test,n,i.i ≤ n → 
-  inject_TM ? (test_char ? test) n i ⊨ 
-  [ tc_true : Rtc_multi_true alpha test n i, Rtc_multi_false alpha test n i ].
-#alpha #test #n #i #Hin #int
-cases (acc_sem_inject … Hin (sem_test_char alpha test) int)
-#k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
-[ @Hloop
-| #Hqtrue lapply (Htrue Hqtrue) * * * #c *
-  #Hcur #Htestc #Hnth_i #Hnth_j %
-  [ %{c} % //
-  | @(eq_vec … (niltape ?)) #i0 #Hi0
-    cases (decidable_eq_nat i0 i) #Hi0i
-    [ >Hi0i @Hnth_i
-    | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
-| #Hqfalse lapply (Hfalse Hqfalse) * * #Htestc #Hnth_i #Hnth_j %
-  [ @Htestc
-  | @(eq_vec … (niltape ?)) #i0 #Hi0
-    cases (decidable_eq_nat i0 i) #Hi0i
-    [ >Hi0i @Hnth_i
-    | @sym_eq @Hnth_j @sym_not_eq // ] ] ]
+lemma right_mk_tape : 
+  ∀sig,ls,c,rs.(c = None ? → ls = [ ] ∨ rs = [ ]) → right ? (mk_tape sig ls c rs) = rs.
+#sig #ls #c #rs cases c // cases ls 
+[ cases rs // 
+| #l0 #ls0 #H normalize cases (H (refl ??)) #H1 [ destruct (H1) | >H1 % ] ]
 qed.
 
-definition Rm_test_null_true ≝ 
-  λalpha,n,i.λt1,t2:Vector ? (S n).
-   current alpha (nth i ? t1 (niltape ?)) ≠ None ? ∧ t2 = t1.
-   
-definition Rm_test_null_false ≝ 
-  λalpha,n,i.λt1,t2:Vector ? (S n).
-    current alpha (nth i ? t1 (niltape ?)) = None ? ∧ t2 = t1.
+lemma left_mk_tape : ∀sig,ls,c,rs.left ? (mk_tape sig ls c rs) = ls.
+#sig #ls #c #rs cases c // cases ls // cases rs //
+qed.
 
-lemma sem_test_null_multi : ∀alpha,n,i.i ≤ n → 
-  inject_TM ? (test_null ?) n i ⊨ 
-    [ tc_true : Rm_test_null_true alpha n i, Rm_test_null_false alpha n i ].
-#alpha #n #i #Hin #int
-cases (acc_sem_inject … Hin (sem_test_null alpha) int)
-#k * #outc * * #Hloop #Htrue #Hfalse %{k} %{outc} % [ %
-[ @Hloop
-| #Hqtrue lapply (Htrue Hqtrue) * * #Hcur #Hnth_i #Hnth_j % //
-  @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) #Hi0i
-  [ >Hi0i @sym_eq @Hnth_i | @sym_eq @Hnth_j @sym_not_eq // ] ] 
-| #Hqfalse lapply (Hfalse Hqfalse) * * #Hcur #Hnth_i #Hnth_j %
-  [ @Hcur
-  | @(eq_vec … (niltape ?)) #i0 #Hi0 cases (decidable_eq_nat i0 i) // 
-    #Hi0i @sym_eq @Hnth_j @sym_not_eq // ] ] 
+lemma current_mk_tape : ∀sig,ls,c,rs.current ? (mk_tape sig ls c rs) = c.
+#sig #ls #c #rs cases c // cases ls // cases rs //
 qed.
 
-definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
-  match (nth src (option sig) v (None ?)) with 
-  [ None ⇒  false 
-  | Some x ⇒  notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
+lemma length_tail : ∀A,l.0 < |l| → |tail A l| < |l|.
+#A * normalize //
+qed.
 
-definition mmove_states ≝ initN 2.
+(*
+[ ]   → [ ], l2, 1
+a::al → 
+      [ ]   → [ ], l1, 2
+      b::bl → match rec(al,bl)
+            x, y, 1 → b::x, y, 1
+            x, y, 2 → a::x, y, 2
+*)
 
-definition mmove0 : mmove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
-definition mmove1 : mmove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
+lemma lists_length_split : 
+ ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
+#A #l1 elim l1
+[ #l2 %{[ ]} %{l2} % % %
+| #hd1 #tl1 #IH *
+  [ %{[ ]} %{(hd1::tl1)} %2 % %
+  | #hd2 #tl2 cases (IH tl2) #x * #y *
+    [ * #IH1 #IH2 %{(hd2::x)} %{y} % normalize % //
+    | * #IH1 #IH2 %{(hd1::x)} %{y} %2 normalize % // ]
+  ]
+]
+qed.
 
-definition trans_mmove ≝ 
- λi,sig,n,D.
- λp:mmove_states × (Vector (option sig) (S n)).
- let 〈q,a〉 ≝ p in match (pi1 … q) with
- [ O ⇒ 〈mmove1,change_vec ? (S n) (null_action ? n) (〈None ?,D〉) i〉
- | S _ ⇒ 〈mmove1,null_action sig n〉 ].
+definition option_cons ≝ λsig.λc:option sig.λl.
+  match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
 
-definition mmove ≝ 
-  λi,sig,n,D.
-  mk_mTM sig n mmove_states (trans_mmove i sig n D) 
-    mmove0 (λq.q == mmove1).
-    
-definition Rm_multi ≝ 
-  λalpha,n,i,D.λt1,t2:Vector ? (S n).
-  t2 = change_vec ? (S n) t1 (tape_move alpha (nth i ? t1 (niltape ?)) D) i.
-   
-lemma sem_move_multi :
-  ∀alpha,n,i,D.i ≤ n → 
-  mmove i alpha n D ⊨ Rm_multi alpha n i D.
-#alpha #n #i #D #Hin #int %{2}
-%{(mk_mconfig ? mmove_states n mmove1 ?)} 
-[| %
- [ whd in ⊢ (??%?); @eq_f whd in ⊢ (??%?); @eq_f %
- | whd >tape_move_multi_def
-   <(change_vec_same … (ctapes …) i (niltape ?))
-   >pmap_change <tape_move_multi_def >tape_move_null_action % ] ]
- qed.
+lemma opt_cons_tail_expand : ∀A,l.l = option_cons A (option_hd ? l) (tail ? l). 
+#A * //
+qed.
+
+definition match_test ≝ λsrc,dst.λsig:DeqSet.λn.λv:Vector ? n.
+  match (nth src (option sig) v (None ?)) with 
+  [ None ⇒  false 
+  | Some x ⇒  notb (nth dst (DeqOption sig) v (None ?) == None ?) ].
   
 definition rewind ≝ λsrc,dst,sig,n.
   parmove src dst sig n L · mmove src sig n R · mmove dst sig n R.
@@ -224,14 +195,14 @@ lemma sem_rewind : ∀src,dst,sig,n.
   src ≠ dst → src < S n → dst < S n → 
   rewind src dst sig n ⊨ R_rewind src dst sig n.
 #src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_rewind_strong …)) //
-#ta #tb * * * #H1 #H2 #H3 #H4 % /2/
+#ta #tb * * * #H1 #H2 #H3 #H4 % /2 by /
 qed.
 
 definition match_step ≝ λsrc,dst,sig,n.
   compare src dst sig n ·
      (ifTM ?? (partest sig n (match_test src dst sig ?))
       (single_finalTM ??
-        (rewind src dst sig n · (inject_TM ? (move_r ?) n dst)))
+        (rewind src dst sig n · mmove dst ?? R))
       (nop …)
       partest1).
 
@@ -292,8 +263,8 @@ lemma sem_match_step :
     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
       (sem_seq … 
         (sem_rewind ???? Hneq Hsrc Hdst) 
-        (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
-      (sem_nop …)))
+        (sem_move_multi … R ?))
+      (sem_nop …))) /2/
 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
@@ -333,15 +304,10 @@ lemma sem_match_step :
           lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
           >change_vec_commute // >change_vec_change_vec
           >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
-          >Hte in Htb; * * #_ >nth_change_vec // #Htb1
-          lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2 %
-          [ @(eq_vec … (niltape ?)) #i #Hi
-            cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
-            [ <(\P Hdsti) >Htb1 >nth_change_vec // >Hmidta_dst
-              >Hrs0 >reverse_reverse cases xs [|#r1 #rs1] %
-            | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)]
-              >Hrs0 >reverse_reverse >nth_change_vec_neq in ⊢ (???%); 
-              <Hrs <Hmidta_src [|@(\Pf Hdsti)] >change_vec_same % ]
+          >Hte in Htb; whd in ⊢ (%→?); #Htb >Htb %
+          [ >change_vec_change_vec >nth_change_vec //
+            >reverse_reverse <Hrs <Hmidta_src >change_vec_same <Hrs0 <Hmidta_dst
+            %
           | >Hmidta_dst %{s'} % [%] #_
             >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
           ]
@@ -349,14 +315,10 @@ lemma sem_match_step :
       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
         [@Htc % % @(not_to_not ??? Hss0) #H destruct (H) %]
         -Htc #Htc destruct (Htc) #_ * #td * whd in ⊢ (%→?); * #_ 
-        #Htd destruct (Htd) * #te * * #_ #Hte * * #_ #Htb1 #Htb2
+        #Htd destruct (Htd) * #te * * #_ #Hte whd in ⊢ (%→?); #Htb
         #s1 #rs1 >Hmidta_src #H destruct (H)
         lapply (Hte … Hmidta_src … Hmidta_dst) -Hte #Hte destruct (Hte) %
-        [ @(eq_vec … (niltape ?)) #i #Hi
-          cases (true_or_false ((dst : DeqNat) == i)) #Hdsti
-          [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
-            cases rs0 [|#r2 #rs2] %
-          | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
+        [ >Htb %
         | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ]
       ]
     ]
@@ -537,12 +499,6 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
 ]
 qed.
 
-axiom daemon : ∀P:Prop.P.
-
-(* XXX: move to turing (or mono) *)
-definition option_cons ≝ λsig.λc:option sig.λl.
-  match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
-
 definition R_match_step_true_naive ≝ 
   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
   |left ? (nth src ? outt (niltape ?))| +
@@ -550,14 +506,6 @@ definition R_match_step_true_naive ≝
   |left ? (nth src ? int (niltape ?))| +
   |option_cons ? (current ? (nth dst ? int (niltape ?))) (right ? (nth dst ? int (niltape ?)))|.
 
-axiom right_mk_tape : ∀sig,ls,c,rs.right ? (mk_tape sig ls c rs) = rs.
-axiom left_mk_tape : ∀sig,ls,c,rs.left ? (mk_tape sig ls c rs) = ls.
-axiom current_mk_tape : ∀sig,ls,c,rs.current ? (mk_tape sig ls c rs) = c.
-axiom length_tail : ∀A,l.0 < |l| → |tail A l| < |l|.
-axiom lists_length_split : 
- ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
-axiom opt_cons_tail_expand : ∀A,l.l = option_cons A (option_hd ? l) (tail ? l).
-  
 lemma sem_match_step_termination :
   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
   match_step src dst sig n ⊨ 
@@ -569,8 +517,8 @@ lemma sem_match_step_termination :
     (acc_sem_if ? n … (sem_partest sig n (match_test src dst sig ?))
       (sem_seq … 
         (sem_rewind_strong ???? Hneq Hsrc Hdst) 
-        (sem_inject … dst (le_S_S_to_le … Hdst) (sem_move_r ? )))
-      (sem_nop …)))
+        (sem_move_multi … R ?))
+      (sem_nop …))) [/2/]
 [ #ta #tb #tc * lapply (refl ? (current ? (nth src ? ta (niltape ?))))
   cases (current ? (nth src ? ta (niltape ?))) in ⊢ (???%→%);
   [ #Hcurta_src #Hcomp #_ * #td * >Hcomp [| % %2 %]
@@ -592,7 +540,7 @@ lemma sem_match_step_termination :
           >nth_current_chars >nth_current_chars >Hcurtc_dst 
           cases (current ? (nth src …)) 
           [normalize in ⊢ (%→?); #H destruct (H)
-          | #x >nth_change_vec // cases (reverse ? rs0)
+          | #x >nth_change_vec [|@Hdst] cases (reverse ? rs0)
             [ normalize in ⊢ (%→?); #H destruct (H)
             | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
@@ -604,8 +552,8 @@ lemma sem_match_step_termination :
           normalize in ⊢ (%→?); #H destruct (H) ]
         | * #xs * #ci * #cj * #rs'' * #rs0' * * * #Hcicj #Hrs #Hrs0
           #Htc * #td * * #Hmatch #Htd destruct (Htd) * #te * * *
-          >Htc >change_vec_commute // >nth_change_vec //
-          >change_vec_commute [|@sym_not_eq //] >nth_change_vec //
+          >Htc >change_vec_commute [|//] >nth_change_vec [|//]
+          >change_vec_commute [|@sym_not_eq //] >nth_change_vec [|//]
           cases (lists_length_split ? ls ls0) #lsa * #lsb * * #Hlen #Hlsalsb
           destruct (Hlsalsb)  *
           [ #Hte #_ #_ <(reverse_reverse … ls) in Hte; <(reverse_reverse … lsa)
@@ -613,14 +561,16 @@ lemma sem_match_step_termination :
             @(list_cases2 … Hlen')
             [ #H1 #H2 >H1 >H2 -H1 -H2 normalize in match (reverse ? [ ]); #Hte #_
               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
-              >change_vec_commute // >change_vec_change_vec
+              >change_vec_commute [|//] >change_vec_change_vec
               >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec #Hte
-              >Hte * * #_ >nth_change_vec // >reverse_reverse 
-              #H lapply (H … (refl ??)) -H #Htb1 #Htb2
-              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)
-              [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
-              >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
-              >right_mk_tape normalize in match (left ??);
+              >Hte whd in ⊢ (%→?); >change_vec_change_vec >nth_change_vec [|//] 
+              >reverse_reverse #Htb
+              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)
+              [ >Htb @eq_f3 // cases (xs@cj::rs0') // ]
+              -Htb #Htb >Htb whd >nth_change_vec [|//]
+              >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec [|//]
+              >right_mk_tape [|cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H)]
+              normalize in match (left ??);
               >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand
               whd in match (option_cons ???); >Hrs0
               normalize in ⊢ (?(?%)%); //
@@ -633,13 +583,20 @@ lemma sem_match_step_termination :
                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
                 >length_reverse >length_reverse destruct (Hlen') //
               | /2 by refl, trans_eq/ ] -Hte
-              #Hte #_ * * #_ >Hte >nth_change_vec // #Htb1 #Htb2
+              #Hte #_  whd in ⊢ (%→?); #Htb
               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
                         (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) 
                         (midtape ? [ ] hdb (reverse sig (reverse sig xs@s0::reverse sig tlb)@ci::rs'')) src)
-              [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
-              >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
-              >right_mk_tape >Hmidta_src >Hmidta_dst 
+              [ >Htb >Hte >nth_change_vec // >change_vec_change_vec >change_vec_commute [|//]
+                >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
+                >change_vec_change_vec >change_vec_commute [|//]
+                @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tla)@cj::rs0') // ]
+              -Htb #Htb >Htb whd 
+              >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
+              >right_mk_tape 
+              [| cases (reverse sig (reverse sig xs@s0::reverse sig tla))
+                 [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
+              >Hmidta_src >Hmidta_dst 
               whd in match (left ??); whd in match (left ??); whd in match (right ??);
               >current_mk_tape <opt_cons_tail_expand whd in match (option_cons ???);
               >Hrs0 >length_append whd in ⊢ (??(??%)); >length_append >length_reverse
@@ -653,13 +610,16 @@ lemma sem_match_step_termination :
               lapply (Hte … (refl ??) … (refl ??) (refl ??)) -Hte
               >change_vec_change_vec >change_vec_commute [|@sym_not_eq //] 
               >change_vec_change_vec #Hte #_
-              >Hte * * #_ >nth_change_vec // >reverse_reverse 
-              #H lapply (H … (refl ??)) -H #Htb1 #Htb2
+              >Hte whd in ⊢ (%→?); >nth_change_vec [|//] >reverse_reverse #Htb 
               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta (mk_tape ? [s0] (option_hd ? (xs@cj::rs0')) (tail ? (xs@cj::rs0'))) dst)
                            (midtape ? lsb s0 (xs@ci::rs'')) src)
-              [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
-              >nth_change_vec_neq // >nth_change_vec //
-              >right_mk_tape normalize in match (left ??);
+              [ >Htb >change_vec_change_vec >change_vec_commute [|//]
+                @eq_f3 // <Hrs0 cases rs0 // ]
+              -Htb #Htb >Htb whd >nth_change_vec [|//]
+              >nth_change_vec_neq [|//] >nth_change_vec [|//]
+              >right_mk_tape 
+              [| cases xs [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ] 
+              normalize in match (left ??);
               >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand >Hrs0
               >length_append normalize >length_append >length_append
               <(reverse_reverse ? lsa) >H1 normalize //
@@ -672,13 +632,20 @@ lemma sem_match_step_termination :
                 >H1 in Hlen'; >H2 whd in ⊢ (??%%→?); #Hlen' 
                 >length_reverse >length_reverse destruct (Hlen') //
               | /2 by refl, trans_eq/ ] -Hte
-              #Hte #_ * * #_ >Hte >nth_change_vec_neq // >nth_change_vec // #Htb1 #Htb2
+              #Hte #_ whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb
               cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
                         (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) 
                         (midtape ? lsb hda (reverse sig (reverse sig xs@s0::reverse sig tla)@ci::rs'')) src)
-              [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
-              >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
-              >right_mk_tape >Hmidta_src >Hmidta_dst 
+              [ >Htb >change_vec_change_vec >change_vec_commute [|//]
+                >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
+                >change_vec_change_vec >change_vec_commute [|//]
+                @eq_f3 // cases (reverse sig (reverse sig xs@s0::reverse sig tlb)@cj::rs0') // ]
+              -Htb #Htb >Htb whd 
+              >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
+              >right_mk_tape
+              [| cases (reverse sig (reverse sig xs@s0::reverse sig tlb))
+                 [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
+              >Hmidta_src >Hmidta_dst 
               whd in match (left ??); whd in match (left ??); whd in match (right ??);
               >current_mk_tape <opt_cons_tail_expand
               whd in match (option_cons ???);
@@ -700,12 +667,15 @@ lemma sem_match_step_termination :
           cut (|reverse ? lsa| = |reverse ? ls|) [ // ] #Hlen'
           @(list_cases2 … Hlen')
           [ #H1 #H2 >H1 >H2 -H1 -H2 #_ #_ normalize in match (reverse ? [ ]); #Hte #_
-            lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte) * * #_
-            >Hmidta_dst #Htb1 lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2
+            lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte)
+            whd in ⊢ (%→?); >Hmidta_dst #Htb
             cut (tb = change_vec ?? ta (mk_tape ? (s0::lsa@lsb) (option_hd ? rs0) (tail ? rs0)) dst)
-            [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
+            [ >Htb cases rs0 // ]
+            -Htb #Htb >Htb whd >nth_change_vec [|//]
             >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst
-            >right_mk_tape normalize in match (left ??); normalize in match (right ??);
+            >right_mk_tape 
+            [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H)] ]
+            normalize in match (left ??); normalize in match (right ??);
             >Hmidta_src >Hmidta_dst >current_mk_tape <opt_cons_tail_expand
             normalize //
           | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
@@ -713,13 +683,17 @@ lemma sem_match_step_termination :
             lapply (Hte ???? (refl ??) ? s0 ? (reverse ? tla) ?? (refl ??))
             [ >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) //
               normalize #H destruct (H) // ] #Hte #_ #_ #_
-            * * #_ >Hte >nth_change_vec // #Htb1 #Htb2
+            whd in ⊢ (%→?); >Hte >change_vec_change_vec >nth_change_vec // #Htb
             cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
                       (mk_tape sig (hda::lsb) (option_hd ? (reverse sig (reverse sig tla)@s0::rs0)) (tail ? (reverse sig (reverse sig tla)@s0::rs0))) dst) 
                       (midtape ? [ ] hdb (reverse sig (reverse sig tlb)@s::rs)) src)
-            [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
-            >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
-            >right_mk_tape >Hmidta_src >Hmidta_dst 
+            [ >Htb >change_vec_commute [|//] @eq_f3 // cases (reverse sig (reverse sig tla)@s0::rs0) // ]
+            -Htb #Htb >Htb whd 
+            >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
+            >right_mk_tape 
+            [| cases (reverse sig (reverse sig tla))
+               [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
+            >Hmidta_src >Hmidta_dst 
             whd in match (left ??); whd in match (left ??); whd in match (right ??);
             >current_mk_tape <opt_cons_tail_expand >length_append
             >length_reverse >length_reverse <(length_reverse ? ls) <Hlen'
@@ -729,11 +703,14 @@ lemma sem_match_step_termination :
          @(list_cases2 … Hlen')
          [ #H1 #H2 >H1 >H2 normalize in match (reverse ? [ ]); #_ #_ #Hte
            lapply (Hte … (refl ??) … (refl ??)) -Hte #Hte destruct (Hte)
-           * * #_ >Hmidta_dst #Htb1 lapply (Htb1 … (refl ??)) -Htb1 #Htb1 #Htb2
+           whd in ⊢ (%→?); #Htb whd >Hmidta_dst
            cut (tb = change_vec (tape sig) (S n) ta (mk_tape ? (s0::ls0) (option_hd ? rs0) (tail ? rs0)) dst)
-           [@daemon] -Htb1 -Htb2 #Htb >Htb whd >nth_change_vec //
+           [ >Htb >Hmidta_dst cases rs0 // ]
+           -Htb #Htb >Htb whd >nth_change_vec [|//]
            >nth_change_vec_neq [|@sym_not_eq //] >Hmidta_src >Hmidta_dst
-           >current_mk_tape >right_mk_tape normalize in ⊢ (??%); <opt_cons_tail_expand
+           >current_mk_tape >right_mk_tape 
+           [| cases rs0 [ #_ %2 % | #x0 #xs0 normalize in ⊢ (??%?→?); #H destruct (H) ]]
+           normalize in ⊢ (??%); <opt_cons_tail_expand
            normalize //
          | #hda #hdb #tla #tlb #H1 #H2 >H1 >H2
            >reverse_cons >reverse_cons #Hte #_ #_
@@ -743,13 +720,17 @@ lemma sem_match_step_termination :
            | >length_reverse >length_reverse cut (|hda::tla| = |hdb::tlb|) //
              normalize #H destruct (H) //
            | /2 by refl, trans_eq/ ] -Hte
-           #Hte * * #_ >Hte >nth_change_vec_neq // >nth_change_vec // #Htb1 #Htb2
+           #Hte whd in ⊢ (%→?); >Hte >nth_change_vec_neq [|//] >nth_change_vec [|//] #Htb
            cut (tb = change_vec ?? (change_vec (tape sig) (S n) ta               
                      (mk_tape sig [hdb] (option_hd ? (reverse sig (reverse sig tlb)@s0::rs0)) (tail ? (reverse sig (reverse sig tlb)@s0::rs0))) dst) 
                      (midtape ? lsb hda (reverse sig (reverse sig tla)@s::rs)) src)
-           [@daemon] -Htb1 -Htb2 #Htb >Htb whd 
-           >nth_change_vec // >nth_change_vec_neq // >nth_change_vec //
-           >right_mk_tape >Hmidta_src >Hmidta_dst 
+           [ >Htb >change_vec_commute [|//] >change_vec_change_vec
+             @eq_f3 // cases (reverse sig (reverse sig tlb)@s0::rs0) // ]
+           -Htb #Htb >Htb whd 
+           >nth_change_vec [|//] >nth_change_vec_neq [|//] >nth_change_vec [|//]
+           >right_mk_tape
+           [| cases (reverse ? (reverse ? tlb)) [|#x0 #xs0] normalize in ⊢ (??%?→?); #H destruct (H) ]
+           >Hmidta_src >Hmidta_dst 
            whd in match (left ??); whd in match (left ??); whd in match (right ??);
            >current_mk_tape <opt_cons_tail_expand >length_append
            normalize in ⊢ (??%); >length_append >reverse_reverse
@@ -788,50 +769,31 @@ lemma sem_match_step_termination :
         normalize #H destruct (H) ] ] ]
 qed.
 
+(* 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. *)
+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 ≝ 
+  match Hwf return (λa0,r.f b = a0 → WF B (λx,y:B. R (f x) (f y)) b) with
+  [ wf a Hwfa ⇒ λHeq.? ] (refl ??).
+% #b1 #HRb @WF_to_WF_f @Hwfa <Heq @HRb
+qed.
+
+lemma lt_WF : ∀n.WF ? lt n.
+#n @(nat_elim1 n) -n #n #IH % @IH
+qed.
 
-definition Pre_match_m ≝ 
-  λsrc,sig,n.λt: Vector (tape sig) (S n).
-  ∃x,xs.
-  nth src (tape sig) t (niltape sig) = midtape ? [] x xs.
-  
 lemma terminate_match_m :
   ∀src,dst,sig,n,t.
   src ≠ dst → src < S n → dst < S n → 
-  Pre_match_m src sig n t → 
   match_m src dst sig n ↓ t.
-#src #dst #sig #n #t #Hneq #Hsrc #Hdst * #start * #xs
-#Hmid_src
-@(terminate_while … (sem_match_step src dst sig n Hneq Hsrc Hdst)) //
-<(change_vec_same … t dst (niltape ?))
-lapply (refl ? (nth dst (tape sig) t (niltape ?))) 
-cases (nth dst (tape sig) t (niltape ?)) in ⊢ (???%→?);
-[ #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
-  >Hmid_src #HR cases (HR ?? (refl ??)) -HR
-  >nth_change_vec // >Htape_dst #_ * #s0 * normalize in ⊢ (%→?); #H destruct (H)
-| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
-  >Hmid_src #HR cases (HR ?? (refl ??)) -HR
-  >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
-  * #s0 * #H destruct (H)
-| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
-  >Hmid_src #HR cases (HR ?? (refl ??)) -HR
-  >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?);
-  * #s0 * #H destruct (H)
-| #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs
-  [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
-   >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ?? (refl ??)) -HR 
-   >change_vec_change_vec #Ht1 #_ % #t2 whd in ⊢ (%→?);
-   >Ht1 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src #HR
-   cases (HR ?? (refl ??)) -HR #_
-   >nth_change_vec // * #s1 * normalize in ⊢ (%→?); #H destruct (H)
-  |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?);
-   >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src
-   #Htrue cases (Htrue ?? (refl ??)) -Htrue >change_vec_change_vec
-   >nth_change_vec // >Hmid_dst whd in match (tape_move_mono ???); #Ht1
-   * #s0 * whd in ⊢ (??%?→?); #H destruct (H) #_ >Ht1
-   lapply (IH t1 ? (s0::ls) r0 ?) 
-   [ >Ht1 >nth_change_vec //
-   | >Ht1 >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
-   | >Ht1 >nth_change_vec // ]
- ]
-]
+#src #dst #sig #n #ta #Hneq #Hsrc #Hdst
+@(terminate_while … (sem_match_step_termination src dst sig n Hneq Hsrc Hdst)) //
+letin f ≝ (λt0:Vector (tape sig) (S n).|left ? (nth src (tape ?) t0 (niltape ?))|
+    +|option_cons ? (current ? (nth dst (tape ?) t0 (niltape ?)))
+      (right ? (nth dst (tape ?) t0 (niltape ?)))|)
+change with (λx,y.f x < f y) in ⊢ (??%?); @WF_to_WF_f @lt_WF
+qed.
+
+lemma sem_match_m : ∀src,dst,sig,n.
+src ≠ dst → src < S n → dst < S n → 
+  match_m src dst sig n \vDash R_match_m src dst sig n.
+#src #dst #sig #n #Hneq #Hsrc #Hdst @WRealize_to_Realize [/2/| @wsem_match_m // ]
 qed.
\ No newline at end of file