]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/moves.ma
cfg_to_obj completed (modulo daemons)
[helm.git] / matita / matita / lib / turing / multi_universal / moves.ma
index 23b872463e62cd96b0d5d0967332315245ff4bb8..f725a1940202b8bf641152ce2eeb21d433e5f98d 100644 (file)
@@ -47,8 +47,8 @@ definition trans_parmove_step ≝
      [ None ⇒ 〈parmove2,null_action ? n〉
      | Some a1 ⇒ 〈parmove1,change_vec ? (S n)
                           (change_vec ?(S n)
-                           (null_action ? n) (Some ? 〈a0,D〉) src)
-                          (Some ? 〈a1,D〉) dst〉 ] ]
+                           (null_action ? n) (〈None sig,D〉) src)
+                          (〈None ?,D〉) dst〉 ] ]
  | S q ⇒ match q with 
    [ O ⇒ (* 1 *) 〈parmove1,null_action ? n〉
    | S _ ⇒ (* 2 *) 〈parmove2,null_action ? n〉 ] ].
@@ -66,8 +66,8 @@ definition R_parmove_step_true ≝
    is_sep x1 = false ∧
    outt = change_vec ?? 
             (change_vec ?? int
-              (tape_move ? (nth src ? int (niltape ?)) (Some ? 〈x1,D〉)) src)
-            (tape_move ? (nth dst ? int (niltape ?)) (Some ? 〈x2,D〉)) dst.
+              (tape_move_mono ? (nth src ? int (niltape ?)) (〈None ?,D〉)) src)
+            (tape_move_mono ? (nth dst ? int (niltape ?)) (〈None ?,D〉)) dst.
 
 definition R_parmove_step_false ≝ 
   λsrc,dst:nat.λsig,n,is_sep.λint,outt: Vector (tape sig) (S n).
@@ -88,7 +88,7 @@ lemma parmove_q0_q2_null_src :
 whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
 @eq_f2
 [ whd in ⊢ (??(???%)?); >Hcurrent %
-| whd in ⊢ (??(???????(???%))?); >Hcurrent @tape_move_null_action ]
+| whd in ⊢ (??(????(???%))?); >Hcurrent @tape_move_null_action ]
 qed.
 
 lemma parmove_q0_q2_sep :
@@ -101,8 +101,8 @@ lemma parmove_q0_q2_sep :
 whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
 @eq_f2
 [ whd in ⊢ (??(???%)?); >Hcurrent whd in ⊢ (??(???%)?); >Hsep %
-| whd in ⊢ (??(???????(???%))?); >Hcurrent
-  whd in ⊢ (??(???????(???%))?); >Hsep @tape_move_null_action ]
+| whd in ⊢ (??(????(???%))?); >Hcurrent
+  whd in ⊢ (??(????(???%))?); >Hsep @tape_move_null_action ]
 qed.
 
 lemma parmove_q0_q2_null_dst :
@@ -116,8 +116,8 @@ lemma parmove_q0_q2_null_dst :
 whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?);
 @eq_f2
 [ whd in ⊢ (??(???%)?); >Hcursrc whd in ⊢ (??(???%)?); >Hsep >Hcurdst %
-| whd in ⊢ (??(???????(???%))?); >Hcursrc
-  whd in ⊢ (??(???????(???%))?); >Hsep >Hcurdst @tape_move_null_action ]
+| whd in ⊢ (??(????(???%))?); >Hcursrc
+  whd in ⊢ (??(????(???%))?); >Hsep >Hcurdst @tape_move_null_action ]
 qed.
 
 lemma parmove_q0_q1 :
@@ -131,19 +131,19 @@ lemma parmove_q0_q1 :
     mk_mconfig ??? parmove1 
      (change_vec ? (S n) 
        (change_vec ?? v
-         (tape_move ? (nth src ? v (niltape ?)) (Some ? 〈a1,D〉)) src)
-       (tape_move ? (nth dst ? v (niltape ?)) (Some ? 〈a2,D〉)) dst).
+         (tape_move_mono ? (nth src ? v (niltape ?)) (〈None ?, D〉)) src)
+       (tape_move_mono ? (nth dst ? v (niltape ?)) (〈None ?, D〉)) dst).
 #src #dst #sig #n #D #is_sep #v #Hneq #Hsrc #Hdst
 #a1 #a2 #Hcursrc #Hcurdst #Hsep
 whd in ⊢ (??%?); >(eq_pair_fst_snd … (trans ????)) whd in ⊢ (??%?); @eq_f2
 [ whd in match (trans ????);
   >Hcursrc >Hcurdst whd in ⊢ (??(???%)?); >Hsep //
 | whd in match (trans ????);
-  >Hcursrc >Hcurdst whd in ⊢ (??(???????(???%))?); >Hsep
-  change with (change_vec ?????) in ⊢ (??(???????%)?);
-  <(change_vec_same … v dst (niltape ?)) in ⊢ (??%?);
-  <(change_vec_same  v src (niltape ?)) in ⊢ (??%?);
-  >pmap_change >pmap_change >tape_move_null_action
+  >Hcursrc >Hcurdst whd in ⊢ (??(????(???%))?); >Hsep whd in ⊢ (??(????(???%))?);
+  <(change_vec_same ?? v dst (niltape ?)) in ⊢ (??%?);
+  >tape_move_multi_def >pmap_change
+  <(change_vec_same ?? v src (niltape ?)) in ⊢ (??%?);
+  >pmap_change <tape_move_multi_def >tape_move_null_action
   @eq_f2 // @eq_f2 // >nth_change_vec_neq //
 ]
 qed.
@@ -159,14 +159,12 @@ cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
 [ #Hcursrc %{2} %
   [| % [ %
     [ whd in ⊢ (??%?); >parmove_q0_q2_null_src /2/
-      <(nth_vec_map ?? (current …) src ? int (niltape ?)) //
     | normalize in ⊢ (%→?); #H destruct (H) ]
     | #_ % // % %2 // ] ]
 | #a #Ha cases (true_or_false (is_sep a)) #Hsep
   [ %{2} %
     [| % [ %
       [ whd in ⊢ (??%?); >(parmove_q0_q2_sep … Hsep) /2/
-        <(nth_vec_map ?? (current …) src ? int (niltape ?)) //
       | normalize in ⊢ (%→?); #H destruct (H) ]
       | #_ % // % % %{a} % // ] ]
   | lapply (refl ? (current ? (nth dst ? int (niltape ?))))
@@ -174,15 +172,11 @@ cases (current ? (nth src ? int (niltape ?))) in ⊢ (???%→?);
     [ #Hcurdst %{2} %
       [| % [ %
         [ whd in ⊢ (??%?); >(parmove_q0_q2_null_dst … Hsep) /2/ 
-          [ <(nth_vec_map ?? (current …) dst ? int (niltape ?)) //
-          | <(nth_vec_map ?? (current …) src ? int (niltape ?)) // ]
         | normalize in ⊢ (%→?); #H destruct (H) ]
         | #_ % // %2 // ] ]
     | #b #Hb %{2} %
       [| % [ % 
         [whd in ⊢  (??%?);  >(parmove_q0_q1 … Hneq Hsrc Hdst ? b ?? Hsep) //
-          [ <(nth_vec_map ?? (current …) dst ? int (niltape ?)) //
-          | <(nth_vec_map ?? (current …) src ? int (niltape ?)) // ]
         | #_ %{a} %{b} % // % // % // ]
         | * #H @False_ind @H % ]
 ]]]]
@@ -233,15 +227,9 @@ lapply (sem_while … (sem_parmove_step src dst sig n L is_sep Hneq Hsrc Hdst) 
       #i #Hi cases (decidable_eq_nat i src) #Hisrc
       [ >Hisrc >(nth_change_vec_neq … src dst) [|@(sym_not_eq … Hneq)]
         >nth_change_vec //
-        >(nth_change_vec_neq … src dst) [|@(sym_not_eq … Hneq)]
-        >nth_change_vec //
       | cases (decidable_eq_nat i dst) #Hidst
-        [ >Hidst >nth_change_vec // >nth_change_vec //
-          >Hdst_tc in Hc1; >Htargetnil
-          normalize in ⊢ (%→?); #Hc1 destruct (Hc1) %
+        [ >Hidst >nth_change_vec // 
         | >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
-          >nth_change_vec_neq [|@(sym_not_eq … Hisrc)]
-          >nth_change_vec_neq [|@(sym_not_eq … Hidst)]
           >nth_change_vec_neq [|@(sym_not_eq … Hisrc)] % 
         ]
       ]
@@ -253,8 +241,7 @@ lapply (sem_while … (sem_parmove_step src dst sig n L is_sep Hneq Hsrc Hdst) 
        >change_vec_change_vec
        >reverse_cons >associative_append
        >reverse_cons >associative_append % 
-     | >Hd >nth_change_vec // >Hdst_tc >Htarget >Hdst_tc in Hc1;
-       normalize in ⊢ (%→?); #H destruct (H) //
+     | >Hd >nth_change_vec //
      | >Hxs in Hlen; >Htarget normalize #Hlen destruct (Hlen) //
      | <Hxs #c1 #Hc1 @Hnosep @memb_cons //
      | >Hd >nth_change_vec_neq [|@sym_not_eq //]