]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/mono.ma
dependences update
[helm.git] / matita / matita / lib / turing / mono.ma
index 88ec04e1c77f1e99d56a291b05693eb53b71dbe5..0fcbfbe480d5d6cc8db8389c8141daa0dca8cde0 100644 (file)
@@ -47,6 +47,19 @@ definition mk_tape :
       | cons r0 rs0 ⇒ leftof ? r0 rs0 ]
     | cons l0 ls0 ⇒ rightof ? l0 ls0 ] ].
 
+lemma current_to_midtape: ∀sig,t,c. current sig t = Some ? c →
+  ∃ls,rs. t = midtape ? ls c rs.
+#sig *
+  [#c whd in ⊢ ((??%?)→?); #Hfalse destruct
+  |#a #l #c whd in ⊢ ((??%?)→?); #Hfalse destruct
+  |#a #l #c whd in ⊢ ((??%?)→?); #Hfalse destruct
+  |#ls #a #rs #c whd in ⊢ ((??%?)→?); #H destruct 
+   @(ex_intro … ls) @(ex_intro … rs) //
+  ]
+qed.
+
+(*********************************** moves ************************************)
+
 inductive move : Type[0] ≝
   | L : move | R : move | N : move.
 
@@ -84,7 +97,18 @@ record config (sig,states:FinSet): Type[0] ≝
 { cstate : states;
   ctape: tape sig
 }.
+
+lemma config_expand: ∀sig,Q,c. 
+  c = mk_config sig Q (cstate ?? c) (ctape ?? c).
+#sig #Q * // 
+qed.
   
+lemma config_eq : ∀sig,M,c1,c2.
+  cstate sig M c1 = cstate sig M c2 → 
+    ctape sig M c1 = ctape sig M c2 →  c1 = c2.
+#sig #M1 * #s1 #t1 * #s2 #t2 //
+qed.
+
 definition step ≝ λsig.λM:TM sig.λc:config sig (states sig M).
   let current_char ≝ current ? (ctape ?? c) in
   let 〈news,mv〉 ≝ trans sig M 〈cstate ?? c,current_char〉 in
@@ -172,10 +196,41 @@ lemma loop_eq : ∀sig,f,q,i,j,a,x,y.
 ]
 qed.
 
+lemma loop_p_true : 
+  ∀A,k,f,p,a.p a = true → loop A (S k) f p a = Some ? a.
+#A #k #f #p #a #Ha normalize >Ha %
+qed.
+
+lemma loop_Some : 
+  ∀A,k,f,p,a,b.loop A k f p a = Some ? b → p b = true.
+#A #k #f #p elim k 
+  [#a #b normalize #Hfalse destruct
+  |#k0 #IH #a #b whd in ⊢ (??%? → ?); cases (true_or_false (p a)) #Hpa
+    [ >Hpa normalize #H1 destruct // | >Hpa normalize @IH ]
+  ]
+qed. 
+
+lemma loop_lift : ∀A,B,k,lift,f,g,h,hlift,c1,c2.
+  (∀x.hlift (lift x) = h x) → 
+  (∀x.h x = false → lift (f x) = g (lift x)) → 
+  loop A k f h c1 = Some ? c2 → 
+  loop B k g hlift (lift c1) = Some ? (lift … c2).
+#A #B #k #lift #f #g #h #hlift #c1 #c2 #Hfg #Hhlift
+generalize in match c1; elim k
+[#c0 normalize in ⊢ (??%? → ?); #Hfalse destruct (Hfalse)
+|#k0 #IH #c0 whd in ⊢ (??%? → ??%?);
+ cases (true_or_false (h c0)) #Hc0 >Hfg >Hc0 normalize
+ [ #Heq destruct (Heq) % | <Hhlift // @IH ]
+qed.
+
 (************************** Realizability *************************************)
 definition loopM ≝ λsig,M,i,cin.
   loop ? i (step sig M) (λc.halt sig M (cstate ?? c)) cin.
 
+lemma loopM_unfold : ∀sig,M,i,cin.
+  loopM sig M i cin = loop ? i (step sig M) (λc.halt sig M (cstate ?? c)) cin.
+// qed.
+
 definition initc ≝ λsig.λM:TM sig.λt.
   mk_config sig (states sig M) (start sig M) t.
 
@@ -215,6 +270,44 @@ definition accRealize ≝ λsig.λM:TM sig.λacc:states sig M.λRtrue,Rfalse.
   loopM sig M i (initc sig M t) = Some ? outc ∧
     (cstate ?? outc = acc → Rtrue t (ctape ?? outc)) ∧ 
     (cstate ?? outc ≠ acc → Rfalse t (ctape ?? outc)).
+    
+notation "M ⊨ [q: R1,R2]" non associative with precedence 45 for @{ 'cmodels $M $q $R1 $R2}.
+interpretation "conditional realizability" 'cmodels M q R1 R2 = (accRealize ? M q R1 R2).
+
+(******************************** monotonicity ********************************)
+lemma Realize_to_Realize : ∀alpha,M,R1,R2.
+  R1 ⊆ R2 → Realize alpha M R1 → Realize alpha M R2.
+#alpha #M #R1 #R2 #Himpl #HR1 #intape
+cases (HR1 intape) -HR1 #k * #outc * #Hloop #HR1
+@(ex_intro ?? k) @(ex_intro ?? outc) % /2/
+qed.
+
+lemma WRealize_to_WRealize: ∀sig,M,R1,R2.
+  R1 ⊆ R2 → WRealize sig M R1 → WRealize ? M R2.
+#alpha #M #R1 #R2 #Hsub #HR1 #intape #i #outc #Hloop
+@Hsub @(HR1 … i) @Hloop
+qed.
+
+lemma acc_Realize_to_acc_Realize: ∀sig,M.∀q:states sig M.∀R1,R2,R3,R4. 
+  R1 ⊆ R3 → R2 ⊆ R4 → M ⊨ [q:R1,R2] → M ⊨ [q:R3,R4].
+#alpha #M #q #R1 #R2 #R3 #R4 #Hsub13 #Hsub24 #HRa #intape
+cases (HRa intape) -HRa #k * #outc * * #Hloop #HRtrue #HRfalse 
+@(ex_intro ?? k) @(ex_intro ?? outc) % 
+  [ % [@Hloop] #Hq @Hsub13 @HRtrue // | #Hq @Hsub24 @HRfalse //]
+qed.
+
+(**************************** A canonical relation ****************************)
+
+definition R_TM ≝ λsig.λM:TM sig.λq.λt1,t2.
+∃i,outc.
+  loopM ? M i (mk_config ?? q t1) = Some ? outc ∧ 
+  t2 = (ctape ?? outc).
+  
+lemma R_TM_to_R: ∀sig,M,R. ∀t1,t2. 
+  M ⊫ R → R_TM ? M (start sig M) t1 t2 → R t1 t2.
+#sig #M #R #t1 #t2 whd in ⊢ (%→?); #HMR * #i * #outc *
+#Hloop #Ht2 >Ht2 @(HMR … Hloop)
+qed.
 
 (******************************** NOP Machine *********************************)
 
@@ -237,6 +330,12 @@ lemma sem_nop :
 @(ex_intro … (mk_config ?? start_nop intape)) % % 
 qed.
 
+lemma nop_single_state: ∀sig.∀q1,q2:states ? (nop sig). q1 = q2.
+normalize #sig * #n #ltn1 * #m #ltm1 
+generalize in match ltn1; generalize in match ltm1;
+<(le_n_O_to_eq … (le_S_S_to_le … ltn1)) <(le_n_O_to_eq … (le_S_S_to_le … ltm1)) 
+// qed.
+
 (************************** Sequential Composition ****************************)
 
 definition seq_trans ≝ λsig. λM1,M2 : TM sig. 
@@ -256,14 +355,9 @@ definition seq ≝ λsig. λM1,M2 : TM sig.
     (λs.match s with 
       [ inl _ ⇒ false | inr s2 ⇒ halt sig M2 s2]). 
 
-notation "a · b" non associative with precedence 65 for @{ 'middot $a $b}.
+notation "a · b" right associative with precedence 65 for @{ 'middot $a $b}.
 interpretation "sequential composition" 'middot a b = (seq ? a b).
 
-definition Rcomp ≝ λA.λR1,R2:relation A.λa1,a2.
-  ∃am.R1 a1 am ∧ R2 am a2.
-  
-interpretation "relation composition" 'compose R1 R2 = (Rcomp ? R1 R2).
-
 definition lift_confL ≝ 
   λsig,S1,S2,c.match c with 
   [ mk_config s t ⇒ mk_config sig (FinSum S1 S2) (inl … s) t ].
@@ -290,7 +384,7 @@ lemma p_halt_liftL : ∀sig,S1,S2,halt,c.
 #sig #S1 #S2 #halt #c cases c #s #t %
 qed.
 
-lemma trans_liftL : ∀sig,M1,M2,s,a,news,move.
+lemma trans_seq_liftL : ∀sig,M1,M2,s,a,news,move.
   halt ? M1 s = false → 
   trans sig M1 〈s,a〉 = 〈news,move〉 → 
   trans sig (seq sig M1 M2) 〈inl … s,a〉 = 〈inl … news,move〉.
@@ -298,7 +392,7 @@ lemma trans_liftL : ∀sig,M1,M2,s,a,news,move.
 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
 qed.
 
-lemma trans_liftR : ∀sig,M1,M2,s,a,news,move.
+lemma trans_seq_liftR : ∀sig,M1,M2,s,a,news,move.
   halt ? M2 s = false → 
   trans sig M2 〈s,a〉 = 〈news,move〉 → 
   trans sig (seq sig M1 M2) 〈inr … s,a〉 = 〈inr … news,move〉.
@@ -306,14 +400,7 @@ lemma trans_liftR : ∀sig,M1,M2,s,a,news,move.
 #Hhalt #Htrans whd in ⊢ (??%?); >Hhalt >Htrans %
 qed.
 
-lemma config_eq : 
-  ∀sig,M,c1,c2.
-  cstate sig M c1 = cstate sig M c2 → 
-  ctape sig M c1 = ctape sig M c2 →  c1 = c2.
-#sig #M1 * #s1 #t1 * #s2 #t2 //
-qed.
-
-lemma step_lift_confR : ∀sig,M1,M2,c0.
+lemma step_seq_liftR : ∀sig,M1,M2,c0.
  halt ? M2 (cstate ?? c0) = false → 
  step sig (seq sig M1 M2) (lift_confR sig (states ? M1) (states ? M2) c0) =
  lift_confR sig (states ? M1) (states ? M2) (step sig M2 c0).
@@ -325,10 +412,10 @@ lemma step_lift_confR : ∀sig,M1,M2,c0.
   | 2,3: #s1 #l1 #Heq #Hhalt 
   |#ls #s1 #rs #Heq #Hhalt ]
   whd in ⊢ (???(????%)); >Heq whd in ⊢ (???%);
-  whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_liftR … Heq) //
+  whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_seq_liftR … Heq) //
 qed.
 
-lemma step_lift_confL : ∀sig,M1,M2,c0.
+lemma step_seq_liftL : ∀sig,M1,M2,c0.
  halt ? M1 (cstate ?? c0) = false → 
  step sig (seq sig M1 M2) (lift_confL sig (states ? M1) (states ? M2) c0) =
  lift_confL sig ?? (step sig M1 c0).
@@ -340,31 +427,9 @@ lemma step_lift_confL : ∀sig,M1,M2,c0.
   | 2,3: #s1 #l1 #Heq #Hhalt 
   |#ls #s1 #rs #Heq #Hhalt ]
   whd in ⊢ (???(????%)); >Heq whd in ⊢ (???%);
-  whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_liftL … Heq) //
+  whd in ⊢ (??(???%)?); whd in ⊢ (??%?); >(trans_seq_liftL … Heq) //
 qed.
 
-lemma loop_lift : ∀A,B,k,lift,f,g,h,hlift,c1,c2.
-  (∀x.hlift (lift x) = h x) → 
-  (∀x.h x = false → lift (f x) = g (lift x)) → 
-  loop A k f h c1 = Some ? c2 → 
-  loop B k g hlift (lift c1) = Some ? (lift … c2).
-#A #B #k #lift #f #g #h #hlift #c1 #c2 #Hfg #Hhlift
-generalize in match c1; elim k
-[#c0 normalize in ⊢ (??%? → ?); #Hfalse destruct (Hfalse)
-|#k0 #IH #c0 whd in ⊢ (??%? → ??%?);
- cases (true_or_false (h c0)) #Hc0 >Hfg >Hc0 normalize
- [ #Heq destruct (Heq) % | <Hhlift // @IH ]
-qed.
-
-lemma loop_Some : 
-  ∀A,k,f,p,a,b.loop A k f p a = Some ? b → p b = true.
-#A #k #f #p elim k 
-  [#a #b normalize #Hfalse destruct
-  |#k0 #IH #a #b whd in ⊢ (??%? → ?); cases (true_or_false (p a)) #Hpa
-    [ >Hpa normalize #H1 destruct // | >Hpa normalize @IH ]
-  ]
-qed. 
-
 lemma trans_liftL_true : ∀sig,M1,M2,s,a.
   halt ? M1 s = true → 
   trans sig (seq sig M1 M2) 〈inl … s,a〉 = 〈inr … (start ? M2),None ?〉.
@@ -396,12 +461,12 @@ cases (HR2 (ctape sig (states ? M1) outc1)) #k2 * #outc2 * #Hloop2 #HM2
   [ * *
    [ #sl #tl whd in ⊢ (??%? → ?); #Hl %
    | #sr #tr whd in ⊢ (??%? → ?); #Hr destruct (Hr) ]
-  || #c0 #Hhalt <step_lift_confL //
+  || #c0 #Hhalt <step_seq_liftL //
   | #x <p_halt_liftL %
   |6:cases outc1 #s1 #t1 %
   |7:@(loop_lift … (initc ?? (ctape … outc1)) … Hloop2) 
     [ * #s2 #t2 %
-    | #c0 #Hhalt <step_lift_confR // ]
+    | #c0 #Hhalt <step_seq_liftR // ]
   |whd in ⊢ (??(???%)?);whd in ⊢ (??%?);
    generalize in match Hloop1; cases outc1 #sc1 #tc1 #Hloop10 
    >(trans_liftL_true sig M1 M2 ??) 
@@ -414,3 +479,10 @@ cases (HR2 (ctape sig (states ? M1) outc1)) #k2 * #outc2 * #Hloop2 #HM2
 ]
 qed.
 
+theorem sem_seq_app: ∀sig.∀M1,M2:TM sig.∀R1,R2,R3.
+  M1 ⊨ R1 → M2 ⊨ R2 → R1 ∘ R2 ⊆ R3 → M1 · M2 ⊨ R3.
+#sig #M1 #M2 #R1 #R2 #R3 #HR1 #HR2 #Hsub
+#t cases (sem_seq … HR1 HR2 t)
+#k * #outc * #Hloop #Houtc @(ex_intro … k) @(ex_intro … outc)
+% [@Hloop |@Hsub @Houtc]
+qed.