]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/unistep_aux.ma
progress
[helm.git] / matita / matita / lib / turing / multi_universal / unistep_aux.ma
index 7d0b63a3a9aeb259301ad45f1877fd8b6c9b06d6..4be6e129f3a7f5f37d4c41312a040093f7a88780 100644 (file)
@@ -59,7 +59,8 @@ definition obj_to_cfg ≝
   mmove cfg FSUnialpha 2 L ·
   (ifTM ?? (inject_TM ? (test_null ?) 2 obj)
     (copy_step obj cfg FSUnialpha 2 ·
-     mmove cfg FSUnialpha 2 L) 
+     mmove cfg FSUnialpha 2 L ·
+     mmove obj FSUnialpha 2 L) 
     (inject_TM ? (write FSUnialpha null) 2 cfg)
      tc_true) ·
   inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
@@ -70,7 +71,7 @@ definition R_obj_to_cfg ≝ λt1,t2:Vector (tape FSUnialpha) 3.
   nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls) (None ?) [ ] → 
   (∀lso,x,rso.nth obj ? t1 (niltape ?) = midtape FSUnialpha lso x rso → 
    t2 = change_vec ?? t1 
-         (mk_tape ? [ ] (option_hd ? (reverse ? (c::ls))) (tail ? (reverse ? (c::ls)))) cfg) ∧
+         (mk_tape ? [ ] (option_hd ? (reverse ? (x::ls))) (tail ? (reverse ? (x::ls)))) cfg) ∧
   (current ? (nth obj ? t1 (niltape ?)) = None ? → 
    t2 = change_vec ?? t1
          (mk_tape ? [ ] (option_hd FSUnialpha (reverse ? (null::ls))) 
@@ -107,7 +108,8 @@ lemma sem_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
     (sem_if ??????????
      (sem_test_null_multi ?? obj ?)
       (sem_seq ?????? (accRealize_to_Realize … (sem_copy_step …))
-       (sem_move_multi ? 2 cfg L ?))
+       (sem_seq ?????? (sem_move_multi ? 2 cfg L ?)
+        (sem_move_multi ? 2 obj L ?)))
       (sem_inject ???? cfg ? (sem_write FSUnialpha null)))
      (sem_seq ?????? (sem_inject ???? cfg ? (sem_move_to_end_l ?))
        (sem_move_multi ? 2 cfg R ?)))) //
@@ -116,16 +118,28 @@ lemma sem_obj_to_cfg : obj_to_cfg ⊨  R_obj_to_cfg.
 #td * *
 [ * #te * * #Hcurtc #Hte
   * destruct (Hte) #te * *
-  [ whd in ⊢ (%→%→?); * #x * #y * * -Hcurtc #Hcurtc1 #Hcurtc2 #Hte #Htd
-    * #tf * * * whd in ⊢ (%→%→%→%→?); #Htf1 #Htf2 #Htf3 #Htb
+  [ whd in ⊢ (%→%→?); * #x * #y * * -Hcurtc #Hcurtc1 #Hcurtc2 #Hte
+    * #tf * whd in ⊢ (%→%→?); #Htf #Htd
+    * #tg * * * whd in ⊢ (%→%→%→%→?); #Htg1 #Htg2 #Htg3 #Htb
     #c #ls #Hta1 %
     [ #lso #x0 #rso #Hta2 >Hta1 in Htc; >eq_mk_tape_rightof 
       whd in match (tape_move ???); #Htc
-      cut (tf = change_vec ?? tc (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg)
-      [@daemon] -Htf1 -Htf2 -Htf3 #Htf destruct (Htf Hte Htd Htc Htb)
-      >change_vec_change_vec >change_vec_change_vec >change_vec_change_vec
-      >nth_change_vec // >tape_move_mk_tape_R
-      @daemon
+      cut (tg = change_vec ?? td (mk_tape ? [ ] (None ?) (reverse ? ls@[x])) cfg)
+      [@daemon] -Htg1 -Htg2 -Htg3 #Htg destruct (Htg Htf Hte Htd Htc Htb)
+      >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 //
+      [ >Hta2 cases rso in Hta2; whd in match (tape_move_mono ???);
+        [ #Hta2 whd in match (tape_move ???); <Hta2 @change_vec_same
+        | #r1 #rs1 #Hta2 whd in match (tape_move ???); <Hta2 @change_vec_same ]
+      | >tape_move_mk_tape_R [| #_ % %] >reverse_cons
+        >nth_change_vec_neq in Hcurtc1; [|@sym_not_eq //] >Hta2
+        normalize in ⊢ (%→?); #H destruct (H) %
+      ]
     | #Hta2 >Htc in Hcurtc1; >nth_change_vec_neq [| @sym_not_eq //]
       >Hta2 #H destruct (H)
     ]
@@ -180,7 +194,8 @@ definition cfg_to_obj ≝
   (ifTM ?? (inject_TM ? test_null_char 2 cfg)
     (nop ? 2)
     (copy_step cfg obj FSUnialpha 2 ·
-     mmove cfg FSUnialpha 2 L) 
+     mmove cfg FSUnialpha 2 L ·
+     mmove obj FSUnialpha 2 L) 
      tc_true) ·
   inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
   mmove cfg FSUnialpha 2 R.
@@ -198,7 +213,7 @@ definition R_cfg_to_obj ≝ λt1,t2:Vector (tape FSUnialpha) 3.
              (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).
 
-axiom sem_cfg_to_obj : obj_to_cfg ⊨  R_obj_to_cfg.
+axiom sem_cfg_to_obj : cfg_to_obj ⊨  R_cfg_to_obj.
 (*@(sem_seq_app FSUnialpha 2 ????? (sem_move_multi ? 2 cfg L ?)
    (sem_seq ??????
     (sem_if ??????????
@@ -252,9 +267,20 @@ 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 *)
-
-axiom tape_move_obj : mTM FSUnialpha 2.
+   
+definition tape_move_obj : mTM FSUnialpha 2 ≝ 
+  ifTM ?? 
+   (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit false)) 2 prg)
+   (mmove obj FSUnialpha 2 L)
+   (ifTM ?? 
+    (inject_TM ? (test_char ? (λc:FSUnialpha.c == bit true)) 2 prg)
+    (mmove obj FSUnialpha 2 R)
+    (nop ??)
+    tc_true)
+   tc_true.
 
 definition unistep ≝ 
-  obj_to_cfg · match_m cfg prg FSUnialpha 2 · copy prg cfg FSUnialpha 2 ·
-   cfg_to_obj · tape_move_obj.
\ No newline at end of file
+  obj_to_cfg · match_m cfg prg FSUnialpha 2 · 
+  inject_TM ? (move_to_end FSUnialpha L) 2 cfg ·
+  mmove cfg FSUnialpha 2 R · copy prg cfg FSUnialpha 2 ·
+  cfg_to_obj · tape_move_obj.
\ No newline at end of file