X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fmono.ma;h=0fcbfbe480d5d6cc8db8389c8141daa0dca8cde0;hb=bd5d6160029247d8c4e3f8cec82f7acd7199d7d5;hp=e559c88dbc7970375f6a6b41bfe50ac62e2d825b;hpb=b0d97cd7e2c50fb1fc2d50c86f3140e226b08a81;p=helm.git diff --git a/matita/matita/lib/turing/mono.ma b/matita/matita/lib/turing/mono.ma index e559c88db..0fcbfbe48 100644 --- a/matita/matita/lib/turing/mono.ma +++ b/matita/matita/lib/turing/mono.ma @@ -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. @@ -261,6 +274,41 @@ definition accRealize ≝ λsig.λM:TM sig.λacc:states sig M.λRtrue,Rfalse. 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 *********************************) (* NO OPERATION @@ -307,7 +355,7 @@ 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 lift_confL ≝