]> matita.cs.unibo.it Git - helm.git/commitdiff
Added a turing/universal directory for the universal turing machine (and
authorWilmer Ricciotti <ricciott@cs.unibo.it>
Fri, 4 May 2012 15:09:54 +0000 (15:09 +0000)
committerWilmer Ricciotti <ricciott@cs.unibo.it>
Fri, 4 May 2012 15:09:54 +0000 (15:09 +0000)
auxiliary definitions).
Added the definition of machine move_char (variant c) to be used in the
universal machine.
Small library refactoring.

matita/matita/lib/basics/deqsets.ma
matita/matita/lib/basics/finset.ma
matita/matita/lib/basics/lists/list.ma
matita/matita/lib/basics/star.ma
matita/matita/lib/turing/mono.ma
matita/matita/lib/turing/while_machine.ma

index 2d6434f6578ffc2f4440aa5843388068be8e22ba..e874eb8ffe6487fd9cd5bc7489c3ed375030ffc3 100644 (file)
@@ -28,6 +28,10 @@ notation "\P H" non associative with precedence 90
 notation "\b H" non associative with precedence 90 
   for @{(proj2 … (eqb_true ???) $H)}. 
   
+notation < "𝐃" non associative with precedence 90 
+ for @{'bigD}.
+interpretation "DeqSet" 'bigD = (mk_DeqSet ???).
+  
 lemma eqb_false: ∀S:DeqSet.∀a,b:S. 
   (eqb ? a b) = false ↔ a ≠ b.
 #S #a #b % #H 
index 3efd26f3f59fcb806a7f795371ba3539a9f7905b..a8ee13a074ad558fbd733c1f2b7718a9d0177c96 100644 (file)
@@ -19,6 +19,10 @@ record FinSet : Type[1] ≝
   enum_unique: uniqueb FinSetcarr enum = true
 }.
 
+notation < "𝐅" non associative with precedence 90 
+ for @{'bigF}.
+interpretation "FinSet" 'bigF = (mk_FinSet ???).
+
 (* bool *)
 lemma bool_enum_unique: uniqueb ? [true;false] = true.
 // qed.
index ff7fcf80f8ee973cbec708aefa8c076365826018..c89bb78560b5ac1b1a9dbb061a372e85318afd9a 100644 (file)
@@ -55,6 +55,11 @@ definition hd ≝ λA.λl: list A.λd:A.
 
 definition tail ≝  λA.λl: list A.
   match l with [ nil ⇒  [] | cons hd tl ⇒  tl].
+  
+definition option_hd ≝ 
+  λA.λl:list A. match l with
+  [ nil ⇒ None ?
+  | cons a _ ⇒ Some ? a ].
 
 interpretation "append" 'append l1 l2 = (append ? l1 l2).
 
index 8fd32c2e1f3092ade0b5a40208b3afb2641c2f54..09730b6f22e586a145bca785f77dba56f30d80c4 100644 (file)
@@ -88,6 +88,34 @@ theorem star_inv: ∀A,R.
 #H (elim H) /2/ normalize #c #d #H1 #H2 #H3 @(trans_star … H3) /2/
 qed.
 
+lemma star_decomp_l : 
+  ∀A,R,x,y.star A R x y → x = y ∨ ∃z.R x z ∧ star A R z y.
+#A #R #x #y #Hstar elim Hstar
+[ #b #c #Hleft #Hright *
+  [ #H1 %2 @(ex_intro ?? c) % //
+  | * #x0 * #H1 #H2 %2 @(ex_intro ?? x0) % /2/ ]
+| /2/ ]
+qed.
+
+axiom star_ind_l : 
+  ∀A:Type[0].∀R:relation A.∀Q:A → A → Prop.
+  (∀a.Q a a) → 
+  (∀a,b,c.R a b → star A R b c → Q b c → Q a c) → 
+  ∀x,y.star A R x y → Q x y.
+(* #A #R #Q #H1 #H2 #x #y #H0 elim H0
+[ #b #c #Hleft #Hright #IH
+  cases (star_decomp_l ???? Hleft)
+  [ #Hx @H2 //
+  | * #z * #H3 #H4 @(H2 … H3) /2/
+[
+|
+generalize in match (λb.H2 x b y); elim H0
+[#b #c #Hleft #Hright #H2' #H3 @H3
+ @(H3 b)
+ elim H0
+[ #b #c #Hleft #Hright #IH //
+| *)
+
 (* RC and star *)
 
 lemma TC_to_star: ∀A,R,a,b.TC A R a b → star A R a b.
index cd9598765f0ba881bf5b1afd3b6d6c6a73d276ad..fae3c7673c5097bbd9998785f3e5679ac0ef7a4b 100644 (file)
@@ -44,6 +44,16 @@ definition current ≝
  λsig.λt:tape sig.match t with
  [ midtape _ c _ ⇒ Some ? c
  | _ ⇒ None ? ].
+definition mk_tape : 
+  ∀sig:FinSet.list sig → option sig → list sig → tape sig ≝ 
+  λsig,lt,c,rt.match c with
+  [ Some c' ⇒ midtape sig lt c' rt
+  | None ⇒ match lt with 
+    [ nil ⇒ match rt with
+      [ nil ⇒ niltape ?
+      | cons r0 rs0 ⇒ leftof ? r0 rs0 ]
+    | cons l0 ls0 ⇒ rightof ? l0 ls0 ] ].
 
 inductive move : Type[0] ≝
 | L : move 
index 2323094b4271588c4b90d3f79a11239daa42090b..52a0a79b75e62d44b7ab1e25abc83af29b7d381c 100644 (file)
@@ -123,261 +123,6 @@ cases (loop_split ?? (λc. halt sig M (cstate ?? c)) ????? Hloop)
  ]
 qed.
 
-(* inductive move_states : Type[0] ≝ 
-| start : move_states
-| q1 : move_states
-| q2 : move_states
-| q3 : move_states
-| qacc : move_states
-| qfail : move_states.
-
-definition 
-*)
-
-definition mystates : FinSet → FinSet ≝ λalpha:FinSet.FinProd (initN 5) alpha.
-
-definition move_char ≝ 
- λalpha:FinSet.λsep:alpha.
- mk_TM alpha (mystates alpha)
- (λp.let 〈q,a〉 ≝ p in
-  let 〈q',b〉 ≝ q in
-  match a with 
-  [ None ⇒ 〈〈4,sep〉,None ?〉 
-  | Some a' ⇒ 
-  match q' with
-  [ O ⇒ (* qinit *)
-    match a' == sep with
-    [ true ⇒ 〈〈4,sep〉,None ?〉
-    | false ⇒ 〈〈1,a'〉,Some ? 〈a',L〉〉 ]
-  | S q' ⇒ match q' with
-    [ O ⇒ (* q1 *)
-      〈〈2,a'〉,Some ? 〈b,R〉〉
-    | S q' ⇒ match q' with
-      [ O ⇒ (* q2 *)
-        〈〈3,sep〉,Some ? 〈b,R〉〉
-      | S q' ⇒ match q' with
-        [ O ⇒ (* qacc *)
-          〈〈3,sep〉,None ?〉
-        | S q' ⇒ (* qfail *)
-          〈〈4,sep〉,None ?〉 ] ] ] ] ])
-  〈0,sep〉
-  (λq.let 〈q',a〉 ≝ q in q' == 3 ∨ q' == 4).
-
-definition mk_tape : 
-  ∀sig:FinSet.list sig → option sig → list sig → tape sig ≝ 
-  λsig,lt,c,rt.match c with
-  [ Some c' ⇒ midtape sig lt c' rt
-  | None ⇒ match lt with 
-    [ nil ⇒ match rt with
-      [ nil ⇒ niltape ?
-      | cons r0 rs0 ⇒ leftof ? r0 rs0 ]
-    | cons l0 ls0 ⇒ rightof ? l0 ls0 ] ].
-    
-lemma cmove_q0_q1 : 
-  ∀alpha:FinSet.∀sep,a,ls,a0,rs.
-  a0 == sep = false → 
-  step alpha (move_char alpha sep)
-    (mk_config ?? 〈0,a〉 (mk_tape … ls (Some ? a0) rs)) =
-  mk_config alpha (states ? (move_char alpha sep)) 〈1,a0〉
-    (tape_move_left alpha ls a0 rs).
-#alpha #sep #a *
-[ #a0 #rs #Ha0 whd in ⊢ (??%?); 
-  normalize in match (trans ???); >Ha0 %
-| #a1 #ls #a0 #rs #Ha0 whd in ⊢ (??%?);
-  normalize in match (trans ???); >Ha0 %
-]
-qed.
-    
-lemma cmove_q1_q2 :
-  ∀alpha:FinSet.∀sep,a,ls,a0,rs.
-  step alpha (move_char alpha sep) 
-    (mk_config ?? 〈1,a〉 (mk_tape … ls (Some ? a0) rs)) = 
-  mk_config alpha (states ? (move_char alpha sep)) 〈2,a0〉 
-    (tape_move_right alpha ls a rs).
-#alpha #sep #a #ls #a0 * //
-qed.
-
-lemma cmove_q2_q3 :
-  ∀alpha:FinSet.∀sep,a,ls,a0,rs.
-  step alpha (move_char alpha sep) 
-    (mk_config ?? 〈2,a〉 (mk_tape … ls (Some ? a0) rs)) = 
-  mk_config alpha (states ? (move_char alpha sep)) 〈3,sep〉 
-    (tape_move_right alpha ls a rs).
-#alpha #sep #a #ls #a0 * //
-qed.
-
-definition option_hd ≝ 
-  λA.λl:list A. match l with
-  [ nil ⇒ None ?
-  | cons a _ ⇒ Some ? a ].
-
-definition Rmove_char_true ≝ 
-  λalpha,sep,t1,t2.
-   ∀a,b,ls,rs. b ≠ sep → 
-    t1 = midtape alpha (a::ls) b rs → 
-    t2 = mk_tape alpha (a::b::ls) (option_hd ? rs) (tail ? rs).
-
-definition Rmove_char_false ≝ 
-  λalpha,sep,t1,t2.
-    left ? t1 ≠ [] → current alpha t1 ≠ None alpha → 
-      current alpha t1 = Some alpha sep ∧ t2 = t1.
-    
-lemma loop_S_true : 
-  ∀A,n,f,p,a.  p a = true → 
-  loop A (S n) f p a = Some ? a. /2/
-qed.
-
-lemma loop_S_false : 
-  ∀A,n,f,p,a.  p a = false → 
-  loop A (S n) f p a = loop A n f p (f a).
-normalize #A #n #f #p #a #Hpa >Hpa %
-qed.
-
-notation < "𝐅" non associative with precedence 90 
- for @{'bigF}.
-notation < "𝐃" non associative with precedence 90 
- for @{'bigD}.
-interpretation "FinSet" 'bigF = (mk_FinSet ???).
-interpretation "DeqSet" 'bigD = (mk_DeqSet ???).
-
-lemma trans_init_sep: 
-  ∀alpha,sep,x.
-  trans ? (move_char alpha sep) 〈〈0,x〉,Some ? sep〉 = 〈〈4,sep〉,None ?〉.
-#alpha #sep #x normalize >(\b ?) //
-qed.
-lemma trans_init_not_sep: 
-  ∀alpha,sep,x,y.y == sep = false → 
-  trans ? (move_char alpha sep) 〈〈0,x〉,Some ? y〉 = 〈〈1,y〉,Some ? 〈y,L〉〉.
-#alpha #sep #x #y #H1 normalize >H1 //
-qed.
-
-lemma sem_move_char :
-  ∀alpha,sep.
-  accRealize alpha (move_char alpha sep) 
-    〈3,sep〉 (Rmove_char_true alpha sep) (Rmove_char_false alpha sep).
-#alpha #sep *
-[@(ex_intro ?? 2)  
-  @(ex_intro … (mk_config ?? 〈4,sep〉 (niltape ?)))
-  % [% [whd in ⊢ (??%?);% |#Hfalse destruct ] |#H1 #H2 @False_ind @(absurd ?? H2) %]
-|#l0 #lt0 @(ex_intro ?? 2)  
-  @(ex_intro … (mk_config ?? 〈4,sep〉 (leftof ? l0 lt0)))
-  % [% [whd in ⊢ (??%?);% |#Hfalse destruct ] |#H1 #H2 @False_ind @(absurd ?? H2) %]
-|#r0 #rt0 @(ex_intro ?? 2)  
-  @(ex_intro … (mk_config ?? 〈4,sep〉 (rightof ? r0 rt0)))
-  % [% [whd in ⊢ (??%?);% |#Hfalse destruct ] |#H1 #H2 #H3 @False_ind @(absurd ?? H3) %]
-| #lt #c #rt cases (true_or_false (c == sep)) #Hc
-  [ @(ex_intro ?? 2) 
-    @(ex_intro ?? (mk_config ?? 〈4,sep〉 (midtape ? lt c rt)))
-    % 
-    [% 
-      [ >(\P Hc) >loop_S_false //
-       >loop_S_true 
-       [ @eq_f whd in ⊢ (??%?); >trans_init_sep %
-       |>(\P Hc) whd in ⊢(??(???(???%))?);
-         >trans_init_sep % ]
-     | #Hfalse destruct
-     ]
-    |#_ #H1 #H2 % // normalize >(\P Hc) % ]
-  | @(ex_intro ?? 4)
-    cases lt
-    [ @ex_intro
-      [|%
-        [ %
-          [ >loop_S_false //
-            >cmove_q0_q1 //
-          | normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
-          ]
-        | normalize in ⊢ (%→?); #_ #H1 @False_ind @(absurd ?? H1) %
-        ]
-      ]
-    | #l0 #lt @ex_intro
-      [| %
-       [ %
-         [ >loop_S_false //
-           >cmove_q0_q1 //
-         | #_ #a #b #ls #rs #Hb #Htape
-           destruct (Htape)
-           >cmove_q1_q2
-           >cmove_q2_q3
-           cases rs normalize //
-         ]
-       | normalize in ⊢ (% → ?); * #Hfalse
-         @False_ind /2/
-       ]
-     ]
-   ]
- ]
-]
-qed.
-
-definition R_while_cmove ≝ 
-  λalpha,sep,t1,t2.
-   ∀a,b,ls,rs,rs'. b ≠ sep → memb ? sep rs = false → 
-    t1 = midtape alpha (a::ls) b (rs@sep::rs') → 
-    t2 = midtape alpha (a::reverse ? rs@b::ls) sep rs'.
-    
-lemma star_cases : 
-  ∀A,R,x,y.star A R x y → x = y ∨ ∃z.R x z ∧ star A R z y.
-#A #R #x #y #Hstar elim Hstar
-[ #b #c #Hleft #Hright *
-  [ #H1 %2 @(ex_intro ?? c) % //
-  | * #x0 * #H1 #H2 %2 @(ex_intro ?? x0) % /2/ ]
-| /2/ ]
-qed.
-
-axiom star_ind_r : 
-  ∀A:Type[0].∀R:relation A.∀Q:A → A → Prop.
-  (∀a.Q a a) → 
-  (∀a,b,c.R a b → star A R b c → Q b c → Q a c) → 
-  ∀x,y.star A R x y → Q x y.
-(* #A #R #Q #H1 #H2 #x #y #H0 elim H0
-[ #b #c #Hleft #Hright #IH
-  cases (star_cases ???? Hleft)
-  [ #Hx @H2 //
-  | * #z * #H3 #H4 @(H2 … H3) /2/
-[
-|
-generalize in match (λb.H2 x b y); elim H0
-[#b #c #Hleft #Hright #H2' #H3 @H3
- @(H3 b)
- elim H0
-[ #b #c #Hleft #Hright #IH //
-| *)
-
-lemma sem_move_char :
-  ∀alpha,sep.
-  WRealize alpha (whileTM alpha (move_char alpha sep) 〈3,sep〉)
-    (R_while_cmove alpha sep).
-#alpha #sep #inc #i #outc #Hloop
-lapply (sem_while … (sem_move_char alpha sep) inc i outc Hloop) [%]
-* #t1 * #Hstar @(star_ind_r ??????? Hstar)
-[ #a whd in ⊢ (% → ?); #H1 #a #b #ls #rs #rs' #Hb #Hrs #Hinc
-  >Hinc in H1; normalize in ⊢ (% → ?); #H1
-  cases (H1 ??)
-  [ #Hfalse @False_ind @(absurd ?? Hb) destruct %
-  |% #H2 destruct (H2)
-  |% #H2 destruct ]
-| #a #b #c #Hstar1 #HRtrue #IH #HRfalse 
-  lapply (IH HRfalse) -IH whd in ⊢ (%→%); #IH
-  #a0 #b0 #ls #rs #rs' #Hb0 #Hrs #Ha
-  whd in Hstar1;
-  lapply (Hstar1 … Hb0 Ha) #Hb
-  @(IH … Hb0 Hrs) >Hb whd in HRfalse;
-  [
-  inc Rtrue* b Rtrue c Rfalse outc
-
-|
-]
-qed.
-  
-   #H1 
-  #a #b #ls #rs #rs #Hb #Hrs #Hinc
-  >Hinc in H2;whd in ⊢ ((??%? → ?) → ?);
-
-lapply (H inc i outc Hloop) *
-
 
 (* (*