]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/match.ma
progress
[helm.git] / matita / matita / lib / turing / multi_universal / match.ma
index da7481f5e1a6494469e3e355d5f8192596bce719..59403fa19a7c7a5e4f591b3ac352ae12d71cbd70 100644 (file)
@@ -192,7 +192,9 @@ definition R_match_step_false ≝
   ((current sig (nth dst (tape sig) int (niltape sig)) = None ?) ∧ outt = int) ∨
     (∃ls0,rs0,xs0. nth dst ? int (niltape ?) = midtape sig ls0 x rs0 ∧
       xs = rs0@xs0 ∧
-      current sig (nth dst (tape sig) outt (niltape sig)) = None ?) ∨
+      outt = change_vec ??
+             (change_vec ?? int (mk_tape sig (reverse ? rs0@x::ls) (option_hd ? xs0) (tail ? xs0)) src)
+             (mk_tape ? (reverse ? rs0@x::ls0) (None ?) [ ]) dst) ∨
     (∃ls0,rs0. 
      nth dst ? int (niltape ?) = midtape sig ls0 x (xs@rs0) ∧
      (* ∀rsj,c. 
@@ -217,14 +219,13 @@ definition R_match_step_true ≝
   ∀s,rs.nth src ? int (niltape ?) = midtape ? [ ] s rs → 
   outt = change_vec ?? int 
          (tape_move_mono … (nth dst ? int (niltape ?)) (〈None ?,R〉)) dst ∧
-  (current sig (nth dst (tape sig) int (niltape sig)) = Some ? s → 
+  (∃s0.current sig (nth dst (tape sig) int (niltape sig)) = Some ? s0 ∧
+   (s0 = s →  
    ∃xs,ci,rs',ls0,cj,rs0.
    rs = xs@ci::rs' ∧
    nth dst ? int (niltape ?) = midtape sig ls0 s (xs@cj::rs0) ∧
-   ci ≠ cj).
+   ci ≠ cj)).
   
-axiom daemon : ∀X:Prop.X.
-      
 lemma sem_match_step :
   ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
   match_step src dst sig n ⊨ 
@@ -257,8 +258,11 @@ lemma sem_match_step :
         #_ #Hcomp cases (Hcomp ????? (refl ??) (refl ??)) -Hcomp [ *
         [ * #rs' * #_ #Hcurtc_dst * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
           >nth_current_chars >nth_current_chars >Hcurtc_dst 
-          cases (current ? (nth src …)) [| #x ] 
-          normalize in ⊢ (%→?); #H destruct (H) 
+          cases (current ? (nth src …)) 
+          [normalize in ⊢ (%→?); #H destruct (H)
+          | #x >nth_change_vec // cases (reverse ? rs0)
+            [ normalize in ⊢ (%→?); #H destruct (H)
+            | #r1 #rs1 normalize in ⊢ (%→?); #H destruct (H) ] ]
         | * #rs0' * #_ #Hcurtc_src * #td * whd in ⊢ (%→?); * whd in ⊢ (??%?→?);
           >(?:nth src ? (current_chars ?? tc) (None ?) = None ?) 
           [|>nth_current_chars >Hcurtc_src >nth_change_vec_neq 
@@ -283,8 +287,8 @@ lemma sem_match_step :
             | <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 % ]
-          | #_ >Hmidta_dst >Hrs0
-            %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
+          | >Hmidta_dst %{s'} % [%] #_
+            >Hrs0 %{xs} %{ci} %{rs''} %{ls0} %{cj} %{rs0'} % // % // 
           ]
         ]
       | lapply (\Pf Hss0) -Hss0 #Hss0 #Htc cut (tc = ta) 
@@ -298,7 +302,7 @@ lemma sem_match_step :
           [ <(\P Hdsti) >(Htb1 … Hmidta_dst) >nth_change_vec // >Hmidta_dst
             cases rs0 [|#r2 #rs2] %
           | <Htb2 [|@(\Pf Hdsti)] >nth_change_vec_neq [| @(\Pf Hdsti)] % ]
-        | >Hs0 #H destruct (H) @False_ind cases (Hss0) /2/ ]
+        | >Hs0 %{s0} % // #H destruct (H) @False_ind cases (Hss0) /2/ ]
       ]
     ]
   ]
@@ -317,7 +321,8 @@ lemma sem_match_step :
       >nth_current_chars >Hta_src >nth_current_chars >Hta_dst 
       whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ] -Hcomp1
       cases (Hcomp2 … Hta_src Hta_dst) [ *
-      [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} % // % //
+      [ * #rs' * #Hxs #Hcurtc % %2 %{ls0} %{rs0} %{rs'} %
+        [ % // | >Hcurtc % ]
       | * #rs0' * #Hxs #Htc %2 >Htc %{ls0} %{rs0'} % // ]
       | * #xs0 * #ci * #cj * #rs' * #rs0' * * *
         #Hci #Hxs #Hrs0 #Htc @False_ind
@@ -338,7 +343,8 @@ definition R_match_m ≝
   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
   ∀x,rs.
   nth src ? int (niltape ?) = midtape sig [ ] x rs →
-  (current sig (nth dst (tape sig) int (niltape sig)) = None ? → outt = int) ∧
+  (current sig (nth dst (tape sig) int (niltape sig)) = None ? → 
+   right ? (nth dst (tape sig) int (niltape sig)) = [ ] → outt = int) ∧
   (∀ls0,x0,rs0.
    nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
    (∃l,l1.x0::rs0 = l@x::rs@l1 ∧
@@ -379,10 +385,12 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
     | * #ls0 * #rs0 * #xs0 * * #Htc_dst #Hrs0 #HNone %
       [ >Htc_dst normalize in ⊢ (%→?); #H destruct (H)
       | #ls1 #x1 #rs1 >Htc_dst #H destruct (H)
-        >Hrs0 cases xs0
+        >Hrs0 >HNone cases xs0
         [ % %{[ ]} %{[ ]} % [ >append_nil >append_nil %]
-          (* change false case 
-             #cj #ls2 #H destruct (H) *) @daemon
+          @eq_f3 //
+          [ >reverse_append %
+          | >reverse_append >reverse_cons >reverse_append
+            >associative_append >associative_append % ]
         | #x2 #xs2 %2 #l #l1 % #Habs lapply (eq_f ?? (length ?) ?? Habs)
           >length_append whd in ⊢ (??%(??%)→?); >length_append
           >length_append normalize >commutative_plus whd in ⊢ (???%→?);
@@ -405,10 +413,18 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
  lapply (refl ? (current ? (nth dst ? ta (niltape ?))))
  cases (current ? (nth dst ? ta (niltape ?))) in ⊢ (???%→?); 
   [#Hcurta_dst % 
-    [#_ whd in Htrue; >Hmidta_src in Htrue; #Htrue
-     cases (Htrue ?? (refl ??)) -Htrue >Hcurta_dst
-     (* dovremmo sapere che ta.dst è sul margine destro, da cui la move non 
-        ha effetto *) #_ cut (tc = ta) [@daemon] #Htc destruct (Htc) #_
+    [#Hcurta_dst #Hrightta_dst whd in Htrue; >Hmidta_src in Htrue; #Htrue
+     cases (Htrue ?? (refl ??)) -Htrue #Htc
+     cut (tc = ta)
+     [ >Htc whd in match (tape_move_mono ???); whd in match (tape_write ???);
+       <(change_vec_same … ta dst (niltape ?)) in ⊢ (???%);
+       lapply Hrightta_dst lapply Hcurta_dst -Hrightta_dst -Hcurta_dst 
+       cases (nth dst ? ta (niltape ?))
+       [ #_ #_ %
+       | #r0 #rs0 #_ normalize in ⊢ (%→?); #H destruct (H)
+       | #l0 #ls0 #_ #_ %
+       | #ls #x0 #rs normalize in ⊢ (%→?); #H destruct (H) ] ]
+     -Htc #Htc destruct (Htc) #_
      cases (IH … Hmidta_src) #Houtc #_ @Houtc //
     |#ls0 #x0 #rs0 #Hmidta_dst >Hmidta_dst in Hcurta_dst;
      normalize in ⊢ (%→?); #H destruct (H)
@@ -418,8 +434,8 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
     #H destruct (H) whd in Htrue; >Hmidta_src in Htrue; #Htrue
     cases (Htrue ?? (refl …)) -Htrue >Hmidta_dst #Htc
     cases (true_or_false (x==c)) #eqx
-    [ lapply (\P eqx) -eqx #eqx destruct (eqx)
-      #Htrue cases (Htrue (refl ??)) -Htrue
+    [ lapply (\P eqx) -eqx #eqx destruct (eqx) * #s0 * whd in ⊢ (??%?→?); #Hs0
+      destruct (Hs0) #Htrue cases (Htrue (refl ??)) -Htrue
       #xs0 * #ci * #rs' * #ls1 * #cj * #rs1 * * #Hxs #H destruct (H) #Hcicj
       >Htc in IH; whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //]
       #IH cases (IH … Hmidta_src) -IH #_ >nth_change_vec //
@@ -428,10 +444,10 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
       #Hxs1 >Hxs1 #IH cases (IH … (refl ??)) -IH
       [ * #l * #l1 * #Hxs1'
         >change_vec_commute // >change_vec_change_vec
-        #Houtc % %{(c::l)} %{l1} % 
+        #Houtc % %{(s0::l)} %{l1} % 
         [ normalize <Hxs1' %
         | >reverse_cons >associative_append >change_vec_commute // @Houtc ]
-      | #H %2 #l #l1 >(?:l@c::xs@l1 = l@(c::xs)@l1) [|%]
+      | #H %2 #l #l1 >(?:l@s0::xs@l1 = l@(s0::xs)@l1) [|%]
         @not_sub_list_merge
         [ #l2 >Hxs <Hxs1 % normalize #H1 lapply (cons_injective_r ????? H1)
           >associative_append #H2 lapply (append_l2_injective ????? (refl ??) H2)
@@ -440,141 +456,75 @@ lapply (sem_while … (sem_match_step src dst sig n Hneq Hsrc Hdst) … Hloop) /
          -H1 #H1 cases (H l2 l3) #H2 @H2 @H1
         ]
       ]
-    | (* in match_step_true manca il caso di fallimento immediato
-         (con i due current diversi) *)
-      @daemon
-      (*
-       #_ lapply (\Pf eqx) -eqx #eqx >Hmidta_dst
-       cases (Htrue ? (refl ??) eqx) -Htrue #Htb #Hendcx #_
-       cases rs0 in Htb;
-       [ #_ %2 #l #l1 cases l
+    | #_ cases (IH x xs ?) -IH
+      [| >Htc >nth_change_vec_neq [|@sym_not_eq //] @Hmidta_src ]
+      >Htc >nth_change_vec // cases rs0
+      [ #_ #_ %2 #l #l1 cases l
        [ normalize cases xs
          [ cases l1
-           [ normalize % #H destruct (H) cases eqx /2/
+           [ normalize % #H destruct (H) cases (\Pf eqx) /2/
            | #tmp1 #l2 normalize % #H destruct (H) ]
          | #tmp1 #l2 normalize % #H destruct (H) ]
        | #tmp1 #l2 normalize % #H destruct (H)cases l2 in e0;
          [ normalize #H1 destruct (H1)
-         | #tmp2 #l3 normalize #H1 destruct (H1) ]
-       ]
-     | #r1 #rs1 normalize in ⊢ (???(????%?)→?); #Htb >Htb in IH; #IH
-       cases (IH ls x xs end rs ? Hnotend Hend Hnotstart) 
-       [| >Htb >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src ] -IH
-       #_ #IH cases (IH Hstart (c::ls0) r1 rs1 ?)
-       [|| >nth_change_vec // ] -IH
-       [ * #l * #l1 * #Hll1 #Hout % %{(c::l)} %{l1} % >Hll1 //
-         >reverse_cons >associative_append #cj0 #ls #Hl1 >(Hout ?? Hl1)
-         >change_vec_commute in ⊢ (??(???%??)?); // @sym_not_eq //
-       | #IH %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@IH]
-         #l1 % #H destruct (H) cases eqx /2/
-       ] *)
+         | #tmp2 #l3 normalize #H1 destruct (H1) ] ]
+      | #r1 #rs1 #_ #IH cases (IH … (refl ??)) -IH
+        [ * #l * #l1 * #Hll1 #Houtc % %{(c::l)} %{l1} % [ >Hll1 % ]
+          >Houtc >change_vec_commute // >change_vec_change_vec 
+          >change_vec_commute [|@sym_not_eq //]
+          >reverse_cons >associative_append %
+        | #Hll1 %2 @(not_sub_list_merge_2 ?? (x::xs)) normalize [|@Hll1]
+         #l1 % #H destruct (H) cases (\Pf eqx) /2/
+        ]
+      ]
     ]
   ]
 ]
 qed.
 
 definition Pre_match_m ≝ 
-  λsrc,sig,n,is_startc,is_endc.λt: Vector (tape sig) (S n).
-  ∃start,xs,end.
-  nth src (tape sig) t (niltape sig) = midtape ? [] start (xs@[end]) ∧
-  is_startc start = true ∧
-  (∀c.c ∈ (xs@[end]) = true → is_startc c = false) ∧
-  (∀c.c ∈ (start::xs) = true → is_endc c = false) ∧
-  is_endc end = true.
+  λ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,is_startc,is_endc,t.
+  ∀src,dst,sig,n,t.
   src ≠ dst → src < S n → dst < S n → 
-  Pre_match_m src sig n is_startc is_endc t → 
-  match_m src dst sig n is_startc is_endc ↓ t.
-#src #dst #sig #n #is_startc #is_endc #t #Hneq #Hsrc #Hdst * #start * #xs * #end
-* * * * #Hmid_src #Hstart #Hnotstart #Hnotend #Hend
-@(terminate_while … (sem_match_step src dst sig n is_startc is_endc Hneq Hsrc Hdst)) //
+  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 normalize in ⊢ (%→?);
-  * #H @False_ind @H %
+  >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 ⊢ (%→?);
-  * #H @False_ind @H %
+  >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 ⊢ (%→?);
-  * #H @False_ind @H %
+  >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 #_
-   #HR cases (HR Hstart Hnotstart)
-   cases (true_or_false (start == s)) #Hs
-   [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
-     cut (∃ci,xs1.xs@[end] = ci::xs1)
-     [ cases xs
-       [ %{end} %{[]} %
-       | #x1 #xs1 %{x1} %{(xs1@[end])} % ] ] * #ci * #xs1 #Hxs
-     >Hxs in Htrue; #Htrue
-     cases (Htrue [ ] start [ ] ? xs1 ? [ ] (refl ??) (refl ??) ?)
-     [ * #_ * #H @False_ind @H % ]
-     #c0 #Hc0 @Hnotend >(memb_single … Hc0) @memb_hd
-   | lapply (\Pf Hs) -Hs #Hs #Htrue #_
-     cases (Htrue ? (refl ??) Hs) -Htrue #Ht1 #_ %
-     #t2 whd in ⊢ (%→?); #HR cases (HR start ?)
-     [ >Ht1 >nth_change_vec // normalize in ⊢ (%→?); * #H @False_ind @H %
-     | >Ht1 >nth_change_vec_neq [|@sym_not_eq //]
-       >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src % ]
-   ]
+   >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 #_ #Htrue
-   <(change_vec_same … t1 dst (niltape ?))
-   cases (Htrue Hstart Hnotstart) -Htrue
-   cases (true_or_false (start == s)) #Hs
-   [ lapply (\P Hs) -Hs #Hs <Hs #_ #Htrue
-    cut (∃ls0,xs0,ci,rs,rs0.
-      nth src ? t (niltape ?) = midtape sig [ ] start (xs0@ci::rs) ∧
-      nth dst ? t (niltape ?) = midtape sig ls0 s (xs0@rs0) ∧
-      (is_endc ci = true ∨ (is_endc ci = false ∧ (∀b,tlb.rs0 = b::tlb → ci ≠ b))))
-    [cases (comp_list ? (xs@[end]) (r0::rs0) is_endc) #xs0 * #xs1 * #xs2
-      * * * #Hxs #Hrs #Hxs0notend #Hcomp >Hrs
-      cut (∃y,ys. xs1 = y::ys)
-      [ lapply Hxs0notend lapply Hxs lapply xs0 elim xs
-        [ *
-          [ normalize #Hxs1 <Hxs1 #_ %{end} %{[]} %
-          | #z #zs normalize in ⊢ (%→?); #H destruct (H) #H
-            lapply (H ? (memb_hd …)) -H >Hend #H1 destruct (H1)
-          ]
-        | #y #ys #IH0 * 
-          [ normalize in ⊢ (%→?); #Hxs1 <Hxs1 #_ %{y} %{(ys@[end])} %
-          | #z #zs normalize in ⊢ (%→?); #H destruct (H) #Hmemb
-            @(IH0 ? e0 ?) #c #Hc @Hmemb @memb_cons // ] ] ] * #y * #ys #Hxs1
-      >Hxs1 in Hxs; #Hxs >Hmid_src >Hmid_dst >Hxs >Hrs
-      %{ls} %{xs0} %{y} %{ys} %{xs2}
-      % [ % // | @Hcomp // ] ]
-    * #ls0 * #xs0 * #ci * #rs * #rs0 * * #Hmid_src' #Hmid_dst' #Hcomp
-    <Hmid_src in Htrue; >nth_change_vec // >Hs #Htrue destruct (Hs)
-    lapply (Htrue ??????? Hmid_src' Hmid_dst' ?) -Htrue
-    [ #c0 #Hc0 @Hnotend cases (orb_true_l … Hc0) -Hc0 #Hc0
-      [ whd in ⊢ (??%?); >Hc0 %
-      | @memb_cons >Hmid_src in Hmid_src'; #Hmid_src' destruct (Hmid_src')
-        lapply e0 -e0 @(list_elim_left … rs)
-        [ #e0 destruct (e0) lapply (append_l1_injective_r ?????? e0) //
-        | #x1 #xs1 #_ >append_cons in ⊢ (???%→?);
-          <associative_append #e0 lapply (append_l1_injective_r ?????? e0) //
-          #e1 >e1 @memb_append_l1 @memb_append_l1 // ] ]
-    | * * #Hciendc cases rs0 in Hcomp;
-      [ #_ * #H @False_ind @H %
-      | #r1 #rs1 * [ >Hciendc #H destruct (H) ]
-        * #_ #Hcomp lapply (Hcomp ?? (refl ??)) -Hcomp #Hcomp #_ #Htrue
-        cases (Htrue ?? (refl ??) Hcomp) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
-        [ >nth_change_vec_neq [|@sym_not_eq //]
-          >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
-        | >nth_change_vec // >Hmid_dst % ] ] ]
-  | >Hmid_dst >nth_change_vec // lapply (\Pf Hs) -Hs #Hs #Htrue #_
-    cases (Htrue ? (refl ??) Hs) #Ht1 #_ >Ht1 @(IH ?? (s::ls) r0)
-    [ >nth_change_vec_neq [|@sym_not_eq //]
-      >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src
-    | >nth_change_vec // ] ] ] ]
+   #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 // ]
+ ]
+]
 qed.
\ No newline at end of file