]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/mono.ma
many changes
[helm.git] / matita / matita / lib / turing / mono.ma
index e670c263569bd8c07f3b78cce937b6de80dc6ab0..393fa2af92784c68e38354c25cdc0d4d2fbaa1b2 100644 (file)
@@ -79,7 +79,53 @@ qed.
 
 inductive move : Type[0] ≝
   | L : move | R : move | N : move.
+  
+(*************************** turning moves into a DeqSet **********************)
+
+definition move_eq ≝ λm1,m2:move.
+  match m1 with
+  [R ⇒ match m2 with [R ⇒ true | _ ⇒ false]
+  |L ⇒ match m2 with [L ⇒ true | _ ⇒ false]
+  |N ⇒ match m2 with [N ⇒ true | _ ⇒ false]].
+
+lemma move_eq_true:∀m1,m2.
+  move_eq m1 m2 = true ↔ m1 = m2.
+*
+  [* normalize [% #_ % |2,3: % #H destruct ]
+  |* normalize [1,3: % #H destruct |% #_ % ]
+  |* normalize [1,2: % #H destruct |% #_ % ]
+qed.
+
+definition DeqMove ≝ mk_DeqSet move move_eq move_eq_true.
+
+unification hint 0 ≔ ;
+    X ≟ DeqMove
+(* ---------------------------------------- *) ⊢ 
+    move ≡ carr X.
+
+unification hint  0 ≔ m1,m2; 
+    X ≟ DeqMove
+(* ---------------------------------------- *) ⊢ 
+    move_eq m1 m2 ≡ eqb X m1 m2.
+
+
+(************************ turning DeqMove into a FinSet ***********************)
+
+definition move_enum ≝ [L;R;N].
+
+lemma move_enum_unique: uniqueb ? [L;R;N] = true.
+// qed.
+
+lemma move_enum_complete: ∀x:move. memb ? x [L;R;N] = true.
+* // qed.
+
+definition FinMove ≝ 
+  mk_FinSet DeqMove [L;R;N] move_enum_unique move_enum_complete.
 
+unification hint  0 ≔ ; 
+    X ≟ FinMove
+(* ---------------------------------------- *) ⊢ 
+    move ≡ FinSetcarr X.
 (********************************** machine ***********************************)
 
 record TM (sig:FinSet): Type[1] ≝