From 75bf98d7d7a16ff8ce2c530e718809e2bc331568 Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Fri, 25 May 2012 08:59:21 +0000 Subject: [PATCH] Porting to the new defintion of finset --- matita/matita/lib/turing/if_machine.ma | 4 +- matita/matita/lib/turing/mono.ma | 6 +- .../matita/lib/turing/universal/alphabet.ma | 12 +- matita/matita/lib/turing/universal/marks.ma | 117 ++++++++++-------- matita/matita/lib/turing/universal/tests.ma | 40 +++--- matita/matita/lib/turing/universal/tuples.ma | 27 ++++ 6 files changed, 133 insertions(+), 73 deletions(-) diff --git a/matita/matita/lib/turing/if_machine.ma b/matita/matita/lib/turing/if_machine.ma index 03d49d4b6..3d6675d02 100644 --- a/matita/matita/lib/turing/if_machine.ma +++ b/matita/matita/lib/turing/if_machine.ma @@ -77,7 +77,7 @@ axiom acc_sem_if: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,acc. accRealize sig M1 acc Rtrue Rfalse → Realize sig M2 R2 → Realize sig M3 R3 → accRealize sig (ifTM sig M1 (single_finalTM … M2) M3 acc) - (inr … (inl … (inr … 0))) + (inr … (inl … (inr … start_nop))) (Rtrue ∘ R2) (Rfalse ∘ R3). @@ -87,7 +87,7 @@ lemma acc_sem_if_app: ∀sig,M1,M2,M3,Rtrue,Rfalse,R2,R3,R4,R5,acc. (∀t1,t2,t3. Rfalse t1 t3 → R3 t3 t2 → R5 t1 t2) → accRealize sig (ifTM sig M1 (single_finalTM … M2) M3 acc) - (inr … (inl … (inr … 0))) + (inr … (inl … (inr … start_nop))) R4 R5. #sig #M1 #M2 #M3 #Rtrue #Rfalse #R2 #R3 #R4 #R5 #acc #HRacc #HRtrue #HRfalse #Hsub1 #Hsub2 diff --git a/matita/matita/lib/turing/mono.ma b/matita/matita/lib/turing/mono.ma index e89d710c9..7d6bad37f 100644 --- a/matita/matita/lib/turing/mono.ma +++ b/matita/matita/lib/turing/mono.ma @@ -290,17 +290,19 @@ definition accRealize ≝ λsig.λM:TM sig.λacc:states sig M.λRtrue,Rfalse:rel *) definition nop_states ≝ initN 1. +definition start_nop : initN 1 ≝ mk_Sig ?? 0 (le_n … (S 0)). definition nop ≝ λalpha:FinSet.mk_TM alpha nop_states (λp.let 〈q,a〉 ≝ p in 〈q,None ?〉) - O (λ_.true). + start_nop (λ_.true). definition R_nop ≝ λalpha.λt1,t2:tape alpha.t2 = t1. lemma sem_nop : ∀alpha.Realize alpha (nop alpha) (R_nop alpha). -#alpha #intape @(ex_intro ?? 1) @ex_intro [| % normalize % ] +#alpha #intape @(ex_intro ?? 1) +@(ex_intro … (mk_config ?? start_nop intape)) % % qed. (* Compositions *) diff --git a/matita/matita/lib/turing/universal/alphabet.ma b/matita/matita/lib/turing/universal/alphabet.ma index 04701aeec..bf2484f29 100644 --- a/matita/matita/lib/turing/universal/alphabet.ma +++ b/matita/matita/lib/turing/universal/alphabet.ma @@ -39,10 +39,18 @@ definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?. |*: * [1,6,11,16: #y ] normalize % #H1 destruct % ] qed. -axiom unialpha_unique : uniqueb DeqUnialpha [bit true;bit false;comma;bar;grid] = true. +lemma unialpha_unique : + uniqueb DeqUnialpha [bit true;bit false;null;comma;bar;grid] = true. +// qed. + +lemma unialpha_complete :∀x:DeqUnialpha. + memb ? x [bit true;bit false;null;comma;bar;grid] = true. +* // * // +qed. definition FSUnialpha ≝ - mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid] unialpha_unique. + mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid] + unialpha_unique unialpha_complete. definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ]. diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index 17992f95f..70000e399 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -30,17 +30,21 @@ include "turing/universal/tests.ma". definition atm_states ≝ initN 3. +definition atm0 : atm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition atm1 : atm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition atm2 : atm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition atmr_step ≝ λalpha:FinSet.λtest:alpha→bool. mk_TM alpha atm_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1, None ?〉 + [ None ⇒ 〈atm1, None ?〉 | Some a' ⇒ match test a' with - [ true ⇒ 〈1,None ?〉 - | false ⇒ 〈2,Some ? 〈a',R〉〉 ]]) - O (λx.notb (x == 0)). + [ true ⇒ 〈atm1,None ?〉 + | false ⇒ 〈atm2,Some ? 〈a',R〉〉 ]]) + atm0 (λx.notb (x == atm0)). definition Ratmr_step_true ≝ λalpha,test,t1,t2. @@ -57,43 +61,45 @@ definition Ratmr_step_false ≝ lemma atmr_q0_q1 : ∀alpha,test,ls,a0,rs. test a0 = true → step alpha (atmr_step alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (atmr_step alpha test)) 1 + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atmr_step alpha test)) atm1 (midtape … ls a0 rs). -#alpha #test #ls #a0 #ts #Htest normalize >Htest % +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest % qed. lemma atmr_q0_q2 : ∀alpha,test,ls,a0,rs. test a0 = false → step alpha (atmr_step alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (atmr_step alpha test)) 2 + (mk_config ?? atm0 (midtape … ls a0 rs)) = + mk_config alpha (states ? (atmr_step alpha test)) atm2 (mk_tape … (a0::ls) (option_hd ? rs) (tail ? rs)). -#alpha #test #ls #a0 #ts #Htest normalize >Htest cases ts // +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest cases ts // qed. lemma sem_atmr_step : ∀alpha,test. accRealize alpha (atmr_step alpha test) - 2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test). + atm2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test). #alpha #test * [ @(ex_intro ?? 2) - @(ex_intro ?? (mk_config ?? 1 (niltape ?))) % - [ % // #Hfalse destruct | #_ % // % % ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (leftof ? a al))) - % [ % // #Hfalse destruct | #_ % // % % ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (rightof ? a al))) - % [ % // #Hfalse destruct | #_ % // % % ] + @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) % + [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al))) + % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ] | #ls #c #rs @(ex_intro ?? 2) cases (true_or_false (test c)) #Htest - [ @(ex_intro ?? (mk_config ?? 1 ?)) + [ @(ex_intro ?? (mk_config ?? atm1 ?)) [| % [ % [ whd in ⊢ (??%?); >atmr_q0_q1 // - | #Hfalse destruct ] + | whd in ⊢ ((??%%)→?); #Hfalse destruct ] | #_ % // %2 @(ex_intro ?? c) % // ] ] - | @(ex_intro ?? (mk_config ?? 2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs)))) + | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs)))) % [ % [ whd in ⊢ (??%?); >atmr_q0_q2 // @@ -116,7 +122,7 @@ definition R_adv_to_mark_r ≝ λalpha,test,t1,t2. t2 = midtape ? (reverse ? rs1@c::ls) b rs2))). definition adv_to_mark_r ≝ - λalpha,test.whileTM alpha (atmr_step alpha test) 2. + λalpha,test.whileTM alpha (atmr_step alpha test) atm2. lemma wsem_adv_to_mark_r : ∀alpha,test. @@ -188,15 +194,18 @@ qed. definition mark_states ≝ initN 2. +definition ms0 : mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)). +definition ms1 : mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)). + definition mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mark_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',true〉,N〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈ms1,None ?〉 + | Some a' ⇒ match (pi1 … q) with + [ O ⇒ let 〈a'',b〉 ≝ a' in 〈ms1,Some ? 〈〈a'',true〉,N〉〉 + | S q ⇒ 〈ms1,None ?〉 ] ]) + ms0 (λq.q == ms1). definition R_mark ≝ λalpha,t1,t2. ∀ls,c,b,rs. @@ -223,16 +232,18 @@ qed. *) definition move_states ≝ initN 2. +definition move0 : move_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)). +definition move1 : move_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)). definition move_r ≝ λalpha:FinSet.mk_TM alpha move_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',R〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈move1,None ?〉 + | Some a' ⇒ match (pi1 … q) with + [ O ⇒ 〈move1,Some ? 〈a',R〉〉 + | S q ⇒ 〈move1,None ?〉 ] ]) + move0 (λq.q == move1). definition R_move_r ≝ λalpha,t1,t2. ∀ls,c,rs. @@ -263,11 +274,11 @@ definition move_l ≝ λalpha:FinSet.mk_TM alpha move_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',L〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈move1,None ?〉 + | Some a' ⇒ match pi1 … q with + [ O ⇒ 〈move1,Some ? 〈a',L〉〉 + | S q ⇒ 〈move1,None ?〉 ] ]) + move0 (λq.q == move1). definition R_move_l ≝ λalpha,t1,t2. ∀ls,c,rs. @@ -297,18 +308,22 @@ qed. definition mrm_states ≝ initN 3. +definition mrm0 : mrm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition mrm1 : mrm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition mrm2 : mrm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition move_right_and_mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mrm_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈2,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ 〈1,Some ? 〈a',R〉〉 + [ None ⇒ 〈mrm2,None ?〉 + | Some a' ⇒ match pi1 … q with + [ O ⇒ 〈mrm1,Some ? 〈a',R〉〉 | S q ⇒ match q with [ O ⇒ let 〈a'',b〉 ≝ a' in - 〈2,Some ? 〈〈a'',true〉,N〉〉 - | S _ ⇒ 〈2,None ?〉 ] ] ]) - O (λq.q == 2). + 〈mrm2,Some ? 〈〈a'',true〉,N〉〉 + | S _ ⇒ 〈mrm2,None ?〉 ] ] ]) + mrm0 (λq.q == mrm2). definition R_move_right_and_mark ≝ λalpha,t1,t2. ∀ls,c,d,b,rs. @@ -337,15 +352,19 @@ qed. definition clear_mark_states ≝ initN 3. +definition clear0 : clear_mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition clear1 : clear_mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition claer2 : clear_mark_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). + definition clear_mark ≝ λalpha:FinSet.mk_TM (FinProd … alpha FinBool) clear_mark_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1,None ?〉 - | Some a' ⇒ match q with - [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',false〉,N〉〉 - | S q ⇒ 〈1,None ?〉 ] ]) - O (λq.q == 1). + [ None ⇒ 〈clear1,None ?〉 + | Some a' ⇒ match pi1 … q with + [ O ⇒ let 〈a'',b〉 ≝ a' in 〈clear1,Some ? 〈〈a'',false〉,N〉〉 + | S q ⇒ 〈clear1,None ?〉 ] ]) + clear0 (λq.q == clear1). definition R_clear_mark ≝ λalpha,t1,t2. ∀ls,c,b,rs. @@ -704,7 +723,7 @@ definition R_comp_step_false ≝ is_marked ? c = false ∧ t2 = t1. lemma sem_comp_step : - accRealize ? comp_step (inr … (inl … (inr … 0))) + accRealize ? comp_step (inr … (inl … (inr … start_nop))) R_comp_step_true R_comp_step_false. #intape cases (acc_sem_if … (sem_test_char ? (is_marked ?)) @@ -769,7 +788,7 @@ cases (acc_sem_if … (sem_test_char ? (is_marked ?)) qed. definition compare ≝ - whileTM ? comp_step (inr … (inl … (inr … 0))). + whileTM ? comp_step (inr … (inl … (inr … start_nop))). (* definition R_compare := diff --git a/matita/matita/lib/turing/universal/tests.ma b/matita/matita/lib/turing/universal/tests.ma index 24a486bc1..2e7092b95 100644 --- a/matita/matita/lib/turing/universal/tests.ma +++ b/matita/matita/lib/turing/universal/tests.ma @@ -22,19 +22,21 @@ include "turing/if_machine.ma". definition tc_states ≝ initN 3. -definition tc_true ≝ 1. +definition tc_start : tc_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)). +definition tc_true : tc_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)). +definition tc_false : tc_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)). definition test_char ≝ λalpha:FinSet.λtest:alpha→bool. mk_TM alpha tc_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈1, None ?〉 + [ None ⇒ 〈tc_true, None ?〉 | Some a' ⇒ match test a' with - [ true ⇒ 〈1,None ?〉 - | false ⇒ 〈2,None ?〉 ]]) - O (λx.notb (x == 0)). + [ true ⇒ 〈tc_true,None ?〉 + | false ⇒ 〈tc_false,None ?〉 ]]) + tc_start (λx.notb (x == tc_start)). definition Rtc_true ≝ λalpha,test,t1,t2. @@ -49,47 +51,49 @@ definition Rtc_false ≝ lemma tc_q0_q1 : ∀alpha,test,ls,a0,rs. test a0 = true → step alpha (test_char alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (test_char alpha test)) 1 + (mk_config ?? tc_start (midtape … ls a0 rs)) = + mk_config alpha (states ? (test_char alpha test)) tc_true (midtape … ls a0 rs). -#alpha #test #ls #a0 #ts #Htest normalize >Htest % +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest % qed. lemma tc_q0_q2 : ∀alpha,test,ls,a0,rs. test a0 = false → step alpha (test_char alpha test) - (mk_config ?? 0 (midtape … ls a0 rs)) = - mk_config alpha (states ? (test_char alpha test)) 2 + (mk_config ?? tc_start (midtape … ls a0 rs)) = + mk_config alpha (states ? (test_char alpha test)) tc_false (midtape … ls a0 rs). -#alpha #test #ls #a0 #ts #Htest normalize >Htest % +#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?); +whd in match (trans … 〈?,?〉); >Htest % qed. lemma sem_test_char : ∀alpha,test. accRealize alpha (test_char alpha test) - 1 (Rtc_true alpha test) (Rtc_false alpha test). + tc_true (Rtc_true alpha test) (Rtc_false alpha test). #alpha #test * [ @(ex_intro ?? 2) - @(ex_intro ?? (mk_config ?? 1 (niltape ?))) % + @(ex_intro ?? (mk_config ?? tc_true (niltape ?))) % [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (leftof ? a al))) +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_true (leftof ? a al))) % [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (rightof ? a al))) +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_true (rightof ? a al))) % [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ] | #ls #c #rs @(ex_intro ?? 2) cases (true_or_false (test c)) #Htest - [ @(ex_intro ?? (mk_config ?? 1 ?)) + [ @(ex_intro ?? (mk_config ?? tc_true ?)) [| % [ % [ whd in ⊢ (??%?); >tc_q0_q1 // | #_ #c0 #Hc0 % // normalize in Hc0; destruct // ] | * #Hfalse @False_ind @Hfalse % ] ] - | @(ex_intro ?? (mk_config ?? 2 (midtape ? ls c rs))) + | @(ex_intro ?? (mk_config ?? tc_false (midtape ? ls c rs))) % [ % [ whd in ⊢ (??%?); >tc_q0_q2 // - | #Hfalse destruct (Hfalse) ] + | whd in ⊢ ((??%%)→?); #Hfalse destruct (Hfalse) ] | #_ #c0 #Hc0 % // normalize in Hc0; destruct (Hc0) // ] ] diff --git a/matita/matita/lib/turing/universal/tuples.ma b/matita/matita/lib/turing/universal/tuples.ma index 909c817e2..f13e1dd73 100644 --- a/matita/matita/lib/turing/universal/tuples.ma +++ b/matita/matita/lib/turing/universal/tuples.ma @@ -51,7 +51,33 @@ qed. definition mk_tuple ≝ λqin,cin,qout,cout,mv. qin @ cin :: 〈comma,false〉:: qout @ cout :: 〈comma,false〉 :: [mv]. + +axiom num_of_state : ∀state:FinSet. state → list (unialpha × bool). + +definition tuple_of_pair ≝ λstates: FinSet. + λhalt:states →bool. + λp: (states × (option FinBool)) × (states × (option (FinBool × move))). + let 〈inp,outp〉 ≝ p in + let 〈q,a〉 ≝ inp in + let cin ≝ match a with [ None ⇒ null | Some b ⇒ bit b ] in + let 〈qn,action〉 ≝ outp in + let 〈cout,mv〉 ≝ + match action with + [ None ⇒ 〈null,null〉 + | Some act ⇒ let 〈na,m〉 ≝ act in + match m with + [ R ⇒ 〈bit na,bit true〉 + | L ⇒ 〈bit na,bit false〉 + | N ⇒ 〈bit na,null〉] + ] in + let qin ≝ num_of_state states q in + let qout ≝ num_of_state states qn in + mk_tuple qin 〈cin,false〉 qout 〈cout,false〉 〈mv,false〉. + + + + (* by definition, a tuple is not marked *) definition tuple_TM : nat → list STape → Prop ≝ λn,t.∃qin,cin,qout,cout,mv. @@ -487,6 +513,7 @@ generalize in match Hc; generalize in match Hl2; cases l2 ] qed. *) + definition init_current ≝ seq ? (adv_to_mark_l ? (is_marked ?)) (seq ? (clear_mark ?) -- 2.39.2