From: Andrea Asperti Date: Tue, 12 Jun 2012 09:46:36 +0000 (+0000) Subject: several changes X-Git-Tag: make_still_working~1640 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=31cb2f0b374657eb5acb95708443e2c1b8481891;p=helm.git several changes --- diff --git a/matita/matita/lib/turing/mono.ma b/matita/matita/lib/turing/mono.ma index e559c88db..8ac107226 100644 --- a/matita/matita/lib/turing/mono.ma +++ b/matita/matita/lib/turing/mono.ma @@ -261,6 +261,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 diff --git a/matita/matita/lib/turing/universal/move_tape.ma b/matita/matita/lib/turing/universal/move_tape.ma index ecc9fa208..ab16d56c7 100644 --- a/matita/matita/lib/turing/universal/move_tape.ma +++ b/matita/matita/lib/turing/universal/move_tape.ma @@ -626,13 +626,6 @@ cases (Hconcrete … Htable ? Ht1) // ] ] ] | #x #Hx @Hbits @memb_append_l1 @Hx ] qed. - -lemma Realize_to_Realize : - ∀alpha,M,R1,R2.(∀t1,t2.R1 t1 t2 → R2 t1 t2) → 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 sem_move_tape_l_abstract : Realize … move_tape_l R_move_tape_l_abstract. @(Realize_to_Realize … mtl_concrete_to_abstract) // diff --git a/matita/matita/lib/turing/universal/normalTM.ma b/matita/matita/lib/turing/universal/normalTM.ma index 686b91099..370c9476e 100644 --- a/matita/matita/lib/turing/universal/normalTM.ma +++ b/matita/matita/lib/turing/universal/normalTM.ma @@ -89,21 +89,21 @@ definition nconfig ≝ λn. config FinBool (initN n). (******************************** tuples **************************************) -(* By general results on FinSets we know that there is every function f between -two finite sets A and B can be described by means of a finite graph of pairs +(* By general results on FinSets we know that every function f between two +finite sets A and B can be described by means of a finite graph of pairs 〈a,f a〉. Hence, the transition function of a normal turing machine can be described by a finite set of tuples 〈i,c〉,〈j,action〉〉 of the following type: (Nat_to n × (option FinBool)) × (Nat_to n × (option (FinBool × move))). -Unfortunately this description is not suitable for an Universal Machine, since +Unfortunately this description is not suitable for a Universal Machine, since such a machine must work with a fixed alphabet, while the size on n is unknown. Hence, we must pass from natural numbers to a representation for them on a finitary, e.g. binary, alphabet. In general, we shall associate -to a pair 〈〈i,c〉,〈j,action〉〉 a tuples with the following syntactical structure +to a pair 〈〈i,c〉,〈j,action〉〉 a tuple with the following syntactical structure |w_ix,w_jy,z where 1. "|" and "," are special characters used as delimiters; 2. w_i and w_j are list of booleans representing the states $i$ and $j$; -3. x is special symbol null if C=None and is a if c=Some a +3. x is special symbol null if c=None and is a if c=Some a 4. y and z are both null if action = None, and are equal to b,m' if action = Some b,m; 5. finally, m' = 0 if m = L, m' = 1 if m=R and m' = null if m = N diff --git a/matita/matita/lib/turing/universal/tuples.ma b/matita/matita/lib/turing/universal/tuples.ma index 6e09b9d52..5f743e2c8 100644 --- a/matita/matita/lib/turing/universal/tuples.ma +++ b/matita/matita/lib/turing/universal/tuples.ma @@ -124,8 +124,8 @@ qed. lemma trans_to_match: ∀n.∀h.∀trans: trans_source n → trans_target n. - ∀inp,outp,qin,cin,qout,cout,mv. trans inp = outp → - tuple_encoding n h 〈inp,outp〉 = mk_tuple qin cin qout cout mv → + ∀s,t,qin,cin,qout,cout,mv. trans s = t → + tuple_encoding n h 〈s,t〉 = mk_tuple qin cin qout cout mv → match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h (graph_enum ?? trans))). #n #h #trans #inp #outp #qin #cin #qout #cout #mv #Htrans #Htuple @(tuple_to_match … (refl…)) Ht1 whd in ⊢ (??%?); @eq_f +normalize in Hqin; destruct (Hqin) % +qed. + definition low_R ≝ λM,qstart,R,t1,t2. ∀tape1. t1 = low_config M (mk_config ?? qstart tape1) → ∃q,tape2.R tape1 tape2 ∧ halt ? M q = true ∧ t2 = low_config M (mk_config ?? q tape2). -definition R_TM ≝ λsig.λM:TM sig.λq.λt1,t2. -∃i,outc. - loop ? i (step sig M) (λc.halt sig M (cstate ?? c)) (mk_config ?? q t1) = Some ? outc ∧ - t2 = (ctape ?? outc). - -(* -definition universal_R ≝ λM,R,t1,t2. - Realize ? M R → - ∀tape1,tape2. - R tape1 tape 2 ∧ - t1 = low_config M (initc ? M tape1) ∧ - ∃q.halt ? M q = true → t2 = low_config M (mk_config ?? q tape2).*) - -axiom uni_step: TM STape. -axiom us_acc: states ? uni_step. - -axiom sem_uni_step: accRealize ? uni_step us_acc low_step_R_true low_step_R_false. +lemma sem_uni_step1: + uni_step ⊨ [us_acc: low_step_R_true, low_step_R_false]. +@(acc_Realize_to_acc_Realize … sem_uni_step) + [@unistep_true_to_low_step | @unistep_false_to_low_step ] +qed. definition universalTM ≝ whileTM ? uni_step us_acc. theorem sem_universal: ∀M:normalTM. ∀qstart. WRealize ? universalTM (low_R M qstart (R_TM FinBool M qstart)). #M #q #intape #i #outc #Hloop -lapply (sem_while … sem_uni_step intape i outc Hloop) +lapply (sem_while … sem_uni_step1 intape i outc Hloop) [@daemon] -Hloop * #ta * #Hstar generalize in match q; -q @(star_ind_l ??????? Hstar) @@ -393,15 +381,6 @@ lapply (sem_while … sem_uni_step intape i outc Hloop) ] qed. -lemma R_TM_to_R: ∀sig,M,R. ∀t1,t2. - WRealize sig M R → R_TM ? M (start ? M) t1 t2 → R t1 t2. -#sig #M #R #t1 #t2 whd in ⊢ (%→?); #HMR * #i * #outc * -#Hloop #Ht2 >Ht2 @(HMR … Hloop) -qed. - -axiom WRealize_to_WRealize: ∀sig,M,R1,R2. - (∀t1,t2.R1 t1 t2 → R2 t1 t2) → WRealize sig M R1 → WRealize ? M R2. - theorem sem_universal2: ∀M:normalTM. ∀R. WRealize ? M R → WRealize ? universalTM (low_R M (start ? M) R). #M #R #HMR lapply (sem_universal … M (start ? M)) @WRealize_to_WRealize