]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/unistep_aux.ma
removed daemons
[helm.git] / matita / matita / lib / turing / multi_universal / unistep_aux.ma
index 9e240013b01c69b6399f56c79c4f3434e074e158..b1f1114092836576b4ace3b161b7c07d2b8b405c 100644 (file)
@@ -88,8 +88,6 @@ lemma eq_mk_tape_rightof :
 #alpha #a #al %
 qed.
 
-axiom daemon : ∀P:Prop.P.
-
 definition option_cons ≝ λsig.λc:option sig.λl.
   match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
 
@@ -103,6 +101,16 @@ lemma tape_move_mk_tape_R :
 normalize //
 qed.
 
+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_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
 @(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?)
    (sem_seq ??????
@@ -126,7 +134,10 @@ lemma sem_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
     [ #lso #x0 #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof 
       whd in match (tape_move ???); #Htc
       cut (tg = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg)
-      [@daemon] -Htg1 -Htg2 -Htg3 #Htg destruct (Htg Htf Hte Htd Htc Htb)
+      [ lapply (eq_vec_change_vec ??????? (Htg2 ls x [ ] ?) Htg3) //
+        >Htd >nth_change_vec_neq // >Htf >nth_change_vec //
+        >Hte >nth_change_vec // >Htc >nth_change_vec // ] -Htg1 -Htg2 -Htg3 #Htg
+      destruct 
       >change_vec_change_vec >change_vec_change_vec
       >change_vec_commute // >change_vec_change_vec
       >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
@@ -158,9 +169,11 @@ lemma sem_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
     >Hta2 normalize in ⊢ (%→?); #H destruct (H)
   | #_ >Hta1 in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
     destruct (Hte) cut (td = change_vec ?? tc (midtape ? ls null []) cfg)
-    [@daemon] -Htd1 -Htd2 #Htd
+    [ lapply (eq_vec_change_vec ??????? (Htd1 ls c [ ] ?) Htd2) // 
+      >Htc >nth_change_vec // ] -Htd1 -Htd2 #Htd
     -Htf1 cut (tf = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[null])) cfg)
-    [@daemon] -Htf2 -Htf3 #Htf destruct (Htf Htd Htc Htb)
+    [ lapply (eq_vec_change_vec ??????? (Htf2 ls null [ ] ?) Htf3) //
+      >Htd >nth_change_vec // ] -Htf2 -Htf3 #Htf destruct (Htf Htd Htc Htb)
     >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
     >change_vec_change_vec >change_vec_change_vec >nth_change_vec //
     >reverse_cons >tape_move_mk_tape_R /2/ ]
@@ -190,11 +203,65 @@ lemma sem_test_null_char :
   | <Houtc % ] ]
 qed.
 
+definition copy_char_states ≝ initN 3.
+
+definition cc0 : copy_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
+definition cc1 : copy_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
+
+definition trans_copy_char ≝ 
+ λsrc,dst.λsig:FinSet.λn.
+ λp:copy_char_states × (Vector (option sig) (S n)).
+ let 〈q,a〉 ≝ p in
+ match pi1 … q with
+ [ O ⇒ 〈cc1,change_vec ? (S n) 
+           (change_vec ? (S n) (null_action ? n) (〈None ?,R〉) src)
+           (〈nth src ? a (None ?),R〉) dst〉
+ | S _ ⇒ 〈cc1,null_action ? n〉 ].
+
+definition copy_char ≝ 
+  λsrc,dst,sig,n.
+  mk_mTM sig n copy_char_states (trans_copy_char src dst sig n) 
+    cc0 (λq.q == cc1).
+
+definition R_copy_char ≝ 
+  λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
+  outt = change_vec ?? 
+         (change_vec ?? int
+          (tape_move_mono ? (nth src ? int (niltape ?)) 〈None ?, R〉) src)
+          (tape_move_mono ? (nth dst ? int (niltape ?)) 
+           〈current ? (nth src ? int (niltape ?)), R〉) dst.
+
+lemma copy_char_q0_q1 :
+  ∀src,dst,sig,n,v.src ≠ dst → src < S n → dst < S n → 
+  step sig n (copy_char src dst sig n) (mk_mconfig ??? cc0 v) =
+    mk_mconfig ??? cc1 
+     (change_vec ? (S n) 
+       (change_vec ?? v
+         (tape_move_mono ? (nth src ? v (niltape ?)) 〈None ?, R〉) src)
+            (tape_move_mono ? (nth dst ? v (niltape ?)) 〈current ? (nth src ? v (niltape ?)), R〉) dst).
+#src #dst #sig #n #v #Heq #Hsrc #Hdst
+whd in ⊢ (??%?);
+<(change_vec_same … v dst (niltape ?)) in ⊢ (??%?);
+<(change_vec_same … v src (niltape ?)) in ⊢ (??%?);
+>tape_move_multi_def @eq_f2 //
+>pmap_change >pmap_change <tape_move_multi_def
+>tape_move_null_action @eq_f2 // @eq_f2
+[ >change_vec_same %
+| >change_vec_same >change_vec_same // ]
+qed.
+
+lemma sem_copy_char:
+  ∀src,dst,sig,n.src ≠ dst → src < S n → dst < S n → 
+  copy_char src dst sig n ⊨ R_copy_char src dst sig n.
+#src #dst #sig #n #Hneq #Hsrc #Hdst #int
+%{2} % [| % [ % | whd >copy_char_q0_q1 // ]]
+qed.
+
 definition cfg_to_obj ≝
   mmove cfg FSUnialpha 2 L ·
   (ifTM ?? (inject_TM ? test_null_char 2 cfg)
     (nop ? 2)
-    (copy_step cfg obj FSUnialpha 2 ·
+    (copy_char cfg obj FSUnialpha 2 ·
      mmove cfg FSUnialpha 2 L ·
      mmove obj FSUnialpha 2 L) 
      tc_true) ·
@@ -213,6 +280,16 @@ definition R_cfg_to_obj ≝ λt1,t2:Vector (tape FSUnialpha) 3.
           (change_vec ?? t1
              (midtape ? (left ? (nth obj ? t1 (niltape ?))) c (right ? (nth obj ? t1 (niltape ?)))) obj)
           (mk_tape ? [ ] (option_hd ? (reverse ? (c::ls))) (tail ? (reverse ? (c::ls)))) cfg).
+          
+lemma tape_move_mk_tape_L :
+  ∀sig,ls,c,rs.
+  (c = None ? → ls = [ ] ∨ rs = [ ]) → 
+  tape_move ? (mk_tape sig ls c rs) L =
+  mk_tape ? (tail ? ls) (option_hd ? ls) (option_cons ? c rs).
+#sig * [ * [ * | #c * ] | #l0 #ls0 * [ *
+[| #r0 #rs0 #H @False_ind cases (H (refl ??)) #H1 destruct (H1) ] | #c * ] ] 
+normalize //
+qed.
 
 lemma sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
 @(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?)
@@ -220,7 +297,7 @@ lemma sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
    (sem_if ??????????
     (acc_sem_inject ?????? cfg ? sem_test_null_char)
     (sem_nop …)
-    (sem_seq ?????? (accRealize_to_Realize … (sem_copy_step …))
+    (sem_seq ?????? (sem_copy_char …)
      (sem_seq ?????? (sem_move_multi ? 2 cfg L ?) (sem_move_multi ? 2 obj L ?))))
    (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?))
     (sem_move_multi ? 2 cfg R ?)))) // [@sym_not_eq //]
@@ -232,9 +309,12 @@ lemma sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
   whd in ⊢ (%→?); #Htb
   #c #ls #Hta %
   [ #Hc >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
-    cut (te = tc) [@daemon] -Hte1 -Hte2 #Hte
+    cut (te = tc)
+    [ lapply (eq_vec_change_vec ??????? (sym_eq … Hte1) Hte2) >change_vec_same // ]
+    -Hte1 -Hte2 #Hte
     cut (tf = change_vec ? 3 te (mk_tape ? [ ] (None ?) (reverse ? ls@[c])) cfg)
-    [@daemon] -Htf1 -Htf2 -Htf3 #Htf
+    [ lapply (eq_vec_change_vec ??????? (Htf2 ls c [ ] ?) Htf3) //
+      >Hte >Htc >nth_change_vec // ] -Htf1 -Htf2 -Htf3 #Htf
     destruct (Htf Hte Htc Htb)
     >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
     >nth_change_vec // >tape_move_mk_tape_R [| #_ % % ] 
@@ -242,22 +322,40 @@ lemma sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
   | #Hc >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
     >Htc in Hcurtc; >nth_change_vec // normalize in ⊢ (%→?); 
     #H destruct (H) @False_ind cases Hc /2/ ]
-| * #te * * * #Hcurtc #Hte1 #Hte2
   * #tf * *
-  [ (* purtroppo copy_step assume che la destinazione sia Some (almeno come semantica) *)
-    STOP
-    * #x * #y * * #Hcurte_cfg #Hcurte_obj #Htf
-    * #tg * whd in ⊢ (%→%→?); #Htg #Htd
-    * #th * * * #Hth1 #Hth2 #Hth3 
-    whd in ⊢ (%→%); #Htb
-    #c #ls #Hta % #Hc
-    [ >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
-      >Htc in Hcurtc; >nth_change_vec // normalize in ⊢ (%→?); >Hc
-      * #H @False_ind /2/
-    | >Hta in Htc; >eq_mk_tape_rightof whd in match (tape_move ???); #Htc
-      cut (te = tc) [@daemon] -Hte1 -Hte2 #Hte
+| * #te * * * #Hcurtc #Hte1 #Hte2
+  * #tf * whd in ⊢ (%→?); #Htf
+  * #tg * whd in ⊢ (%→%→?); #Htg #Htd
+  * #th * * * #Hth1 #Hth2 #Hth3
+  whd in ⊢ (%→?); #Htb 
+  #c #ls #Hta % #Hc
+  [ >Htc in Hcurtc; >Hta >nth_change_vec // >tape_move_mk_tape_L //
+    >Hc normalize in ⊢ (%→?); * #H @False_ind /2/
+  | cut (te = tc)
+    [ lapply (eq_vec_change_vec ??????? (sym_eq … Hte1) Hte2)
+      >change_vec_same // ] -Hte1 -Hte2 #Hte
+    cut (th = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[c])) cfg)
+    [ lapply (eq_vec_change_vec ??????? (Hth2 ls c [ ] ?) Hth3) //
+      >Htd >nth_change_vec_neq // >Htg >nth_change_vec //
+      >Htf >nth_change_vec_neq // >nth_change_vec // 
+      >Hte >Htc >nth_change_vec // >Hta // ] -Hth1 -Hth2 -Hth3 #Hth
+    destruct (Hth Hte Hta Htb Htd Htg Htc Htf) 
+    >change_vec_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 // >change_vec_change_vec    
+    >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //] 
+    >nth_change_vec // >nth_change_vec_neq [|@sym_not_eq //]
+    >change_vec_commute [|@sym_not_eq //]
+    @eq_f3 //
+    [ >Hta >tape_move_mk_tape_L // >nth_change_vec // whd in match (current ??);
+      @eq_f2 // cases (nth obj ? ta (niltape ?))
+      [| #r0 #rs0 | #l0 #ls0 | #ls0 #c0 #rs0 ] try %
+      cases rs0 //
+    | >reverse_cons >tape_move_mk_tape_R // #_ % % ]
+  ]
+]
 qed.
-*)
 
 (* macchina che muove il nastro obj a destra o sinistra a seconda del valore
    del current di prg, che codifica la direzione in cui ci muoviamo *)
@@ -281,6 +379,68 @@ definition tape_move_obj : mTM FSUnialpha 2 ≝
     tc_true)
    tc_true.
 
+definition R_tape_move_obj' ≝ λt1,t2:Vector (tape FSUnialpha) 3.
+  (current ? (nth prg ? t1 (niltape ?)) = Some ? (bit false) → 
+   t2 = change_vec ?? t1 (tape_move ? (nth obj ? t1 (niltape ?)) L) obj) ∧
+  (current ? (nth prg ? t1 (niltape ?)) = Some ? (bit true) → 
+   t2 = change_vec ?? t1 (tape_move ? (nth obj ? t1 (niltape ?)) R) obj) ∧
+  (current ? (nth prg ? t1 (niltape ?)) ≠ Some ? (bit false) →
+   current ? (nth prg ? t1 (niltape ?)) ≠ Some ? (bit true) →  
+   t2 = t1).
+   
+lemma sem_tape_move_obj' : tape_move_obj ⊨ R_tape_move_obj'.
+#ta cases (sem_if ??????????
+  (acc_sem_inject ?????? prg ? (sem_test_char ? (λc:FSUnialpha.c == bit false)))
+  (sem_move_multi ? 2 obj L ?)
+  (sem_if ??????????
+   (acc_sem_inject ?????? prg ? (sem_test_char ? (λc:FSUnialpha.c == bit true)))
+   (sem_move_multi ? 2 obj R ?)
+   (sem_nop …)) ta) //
+#i * #outc * #Hloop #HR %{i} %{outc} % [@Hloop] -i
+cases HR -HR
+[ * #tb * * * * #c * #Hcurta_prg #Hc lapply (\P Hc) -Hc #Hc #Htb1 #Htb2
+  whd in ⊢ (%→%); #Houtc >Houtc -Houtc % [ %
+  [ >Hcurta_prg #H destruct (H) >(?:tb = ta) 
+    [| lapply (eq_vec_change_vec ??????? Htb1 Htb2)
+       >change_vec_same // ] %
+  | >Hcurta_prg #H destruct (H) destruct (Hc) ]
+  | >Hcurta_prg >Hc * #H @False_ind /2/ ]
+| * #tb * * * #Hnotfalse #Htb1 #Htb2 cut (tb = ta) 
+  [ lapply (eq_vec_change_vec ??????? Htb1 Htb2)
+     >change_vec_same // ] -Htb1 -Htb2 #Htb destruct (Htb) *
+  [ * #tc * * * * #c * #Hcurta_prg #Hc lapply (\P Hc) -Hc #Hc #Htc1 #Htc2
+    whd in ⊢ (%→%); #Houtc >Houtc -Houtc % [ %
+    [ >Hcurta_prg #H destruct (H) destruct (Hc)
+    | >Hcurta_prg #H destruct (H) >(?:tc = ta) 
+      [| lapply (eq_vec_change_vec ??????? Htc1 Htc2)
+        >change_vec_same // ] % ]
+    | >Hcurta_prg >Hc #_ * #H @False_ind /2/ ]
+  | * #tc * * * #Hnottrue #Htc1 #Htc2 cut (tc = ta) 
+    [ lapply (eq_vec_change_vec ??????? Htc1 Htc2)
+      >change_vec_same // ] -Htc1 -Htc2 
+    #Htc destruct (Htc) whd in ⊢ (%→?); #Houtc % [ %
+    [ #Hcurta_prg lapply (\Pf (Hnotfalse ? Hcurta_prg)) * #H @False_ind /2/
+    | #Hcurta_prg lapply (\Pf (Hnottrue ? Hcurta_prg)) * #H @False_ind /2/ ]
+    | #_ #_ @Houtc ]
+  ]
+]
+qed.
+
+definition R_tape_move_obj ≝ λt1,t2:Vector (tape FSUnialpha) 3.
+  ∀c. current ? (nth prg ? t1 (niltape ?)) = Some ? c → 
+  t2 = change_vec ?? t1 (tape_move ? (nth obj ? t1 (niltape ?)) (char_to_move c)) obj.
+
+lemma sem_tape_move_obj : tape_move_obj ⊨ R_tape_move_obj.
+@(Realize_to_Realize … sem_tape_move_obj')
+#ta #tb * * #Htb1 #Htb2 #Htb3 * [ *
+[ @Htb2 | @Htb1 ] 
+| #Hcurta_prg change with (nth obj ? ta (niltape ?)) in match (tape_move ???);
+  >change_vec_same @Htb3 >Hcurta_prg % #H destruct (H)
+| #Hcurta_prg change with (nth obj ? ta (niltape ?)) in match (tape_move ???);
+  >change_vec_same @Htb3 >Hcurta_prg % #H destruct (H)
+]
+qed.
+
 definition restart_tape ≝ λi. 
   inject_TM ? (move_to_end FSUnialpha L) 2 i ·
   mmove i FSUnialpha 2 R. 
@@ -354,13 +514,14 @@ lemma low_char_current : ∀t.
   = low_char (current FinBool t).
 * // qed.
 
-definition low_tapes ≝ λM:normalTM.λc:nconfig (no_states M).Vector_of_list ?
+definition low_tapes: ∀M:normalTM.∀c:nconfig (no_states M).Vector ? 3 ≝ 
+λM:normalTM.λc:nconfig (no_states M).Vector_of_list ?
   [tape_map ?? bit (ctape ?? c);
    midtape ? [ ] bar 
     ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]);
    midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
   ].
-  
+
 lemma obj_low_tapes: ∀M,c.
   nth obj ? (low_tapes M c) (niltape ?) = tape_map ?? bit (ctape ?? c).
 // qed.
@@ -407,14 +568,15 @@ lemma map_move_mono: ∀t,cout,m.
 @map_action
 qed. 
 
-definition R_unistep_high ≝ λM:normalTM.λc:nconfig (no_states M).λt1,t2.
+definition R_unistep_high ≝ λM:normalTM.λt1,t2.
+∀c:nconfig (no_states M).
   t1 = low_tapes M c → 
   t2 = low_tapes M (step ? M c). 
 
-lemma R_unistep_equiv : ∀M,c,t1,t2. 
+lemma R_unistep_equiv : ∀M,t1,t2. 
   R_unistep (no_states M) (graph_enum ?? (ntrans M)) (nhalt M) t1 t2 →
-  R_unistep_high M t1 t2.
-#M #c #t1 #t2 #H #Ht1
+  R_unistep_high M t1 t2.
+#M #t1 #t2 #H whd whd in match (nconfig ?); #c #Ht1
 lapply (initial_bar ? (nhalt M) (graph_enum ?? (ntrans M)) (nTM_nog ?)) #Htable
 (* tup = current tuple *)
 cut (∃t.t = 〈〈cstate … c,current ? (ctape … c)〉, 
@@ -456,7 +618,7 @@ lapply(H (bits_of_state ? (nhalt M) (cstate ?? c))
 lapply(H (bits_of_state … (nhalt M) qout) (low_char … cout) 
          (low_mv … m) tup ? Hingraph)
   [>Htup whd in ⊢ (??%?); @eq_f >associative_append %] -H
-#Ht2 >Ht2 @(eq_vec ? 3 … (niltape ?)) #i #Hi 
+#Ht2 @(eq_vec ? 3 ?? (niltape ?) ?) >Ht2 #i #Hi 
 cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
   [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
     [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
@@ -476,7 +638,3 @@ cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
    >Ht1 >prg_low_tapes //
   ]
 qed. 
-   
-
-    
-         
\ No newline at end of file