X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Ftuples.ma;h=9207b8582fafb0a47f0bf195a91ab512ad74a264;hb=3447747453bbf439d071d21dcb68149cae3a9068;hp=2baf597e7ee7e8bc8d02dbc71afb458a657d8c7e;hpb=0a8212f3e87b75e8ab47dc853e612a9a3e1d2544;p=helm.git diff --git a/matita/matita/lib/turing/universal/tuples.ma b/matita/matita/lib/turing/universal/tuples.ma index 2baf597e7..9207b8582 100644 --- a/matita/matita/lib/turing/universal/tuples.ma +++ b/matita/matita/lib/turing/universal/tuples.ma @@ -10,234 +10,284 @@ V_____________________________________________________________*) -(* COMPARE BIT +(****************************** table of tuples *******************************) +include "turing/universal/normalTM.ma". -*) +(* a well formed table is a list of tuples *) + +inductive table_TM (n:nat) : list STape → Prop ≝ +| ttm_nil : table_TM n [] +| ttm_cons : ∀t1,T.tuple_TM n t1 → table_TM n T → table_TM n (t1@T). -include "turing/universal/marks.ma". +lemma wftable: ∀n,h,l.table_TM (S n) (flatten ? (tuples_list n h l)). +#n #h #l elim l // -l #a #tl #Hind +whd in match (flatten … (tuples_list …)); +@ttm_cons // +qed. -definition STape ≝ FinProd … FSUnialpha FinBool. +(*********************** general properties of tables *************************) +lemma no_grids_in_table: ∀n.∀l.table_TM n l → no_grids l. +#n #l #t elim t + [normalize #c #H destruct + |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_grids_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx + cases (memb_append … Hx) -Hx #Hx + [ @(Ht1 … Hx) + | @(IH … Hx) ] ] +qed. -definition only_bits ≝ λl. - ∀c.memb STape c l = true → is_bit (\fst c) = true. - -definition no_grids ≝ λl. - ∀c.memb STape c l = true → is_grid (\fst c) = false. +lemma no_marks_in_table: ∀n.∀l.table_TM n l → no_marks l. +#n #l #t elim t + [normalize #c #H destruct + |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_marks_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx + cases (memb_append … Hx) -Hx #Hx + [ @(Ht1 … Hx) + | @(IH … Hx) ] ] +qed. -definition no_bars ≝ λl. - ∀c.memb STape c l = true → is_bar (\fst c) = false. +axiom last_of_table: ∀n,l,b.¬ table_TM n (l@[〈bar,b〉]). -definition no_marks ≝ λl. - ∀c.memb STape c l = true → is_marked ? c = false. - -definition tuple_TM : nat → list STape → Prop ≝ - λn,t.∃qin,qout,mv,b1,b2. - only_bits qin ∧ only_bits qout ∧ only_bits mv ∧ - |qin| = n ∧ |qout| = n (* ∧ |mv| = ? *) ∧ - t = qin@〈comma,b1〉::qout@〈comma,b2〉::mv. - -inductive table_TM : nat → list STape → Prop ≝ -| ttm_nil : ∀n.table_TM n [] -| ttm_cons : ∀n,t1,T,b.tuple_TM n t1 → table_TM n T → table_TM n (t1@〈bar,b〉::T). +(************************** matching in a table *******************************) +inductive match_in_table (n:nat) (qin:list STape) (cin: STape) + (qout:list STape) (cout:STape) (mv:STape) +: list STape → Prop ≝ +| mit_hd : + ∀tb. + tuple_TM n (mk_tuple qin cin qout cout mv) → + match_in_table n qin cin qout cout mv + (mk_tuple qin cin qout cout mv @tb) +| mit_tl : + ∀qin0,cin0,qout0,cout0,mv0,tb. + tuple_TM n (mk_tuple qin0 cin0 qout0 cout0 mv0) → + match_in_table n qin cin qout cout mv tb → + match_in_table n qin cin qout cout mv + (mk_tuple qin0 cin0 qout0 cout0 mv0@tb). -(* -l0 x* a l1 x0* a0 l2 ------> l0 x a* l1 x0 a0* l2 - ^ ^ +lemma tuple_to_match: ∀n,h,l,qin,cin,qout,cout,mv,p. + p = mk_tuple qin cin qout cout mv + → mem ? p (tuples_list n h l) → + match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h l)). +#n #h #l #qin #cin #qout #cout #mv #p +#Hp elim l + [whd in ⊢ (%→?); @False_ind + |#p1 #tl #Hind * + [#H whd in match (tuples_list ???); + Hp @mit_hd // + |#H whd in match (tuples_list ???); + cases (is_tuple n h p1) #qin1 * #cin1 * #qout1 * #cout1 * #mv1 + * #_ #Htuplep1 >Htuplep1 @mit_tl // @Hind // + ] + ] +qed. + +axiom match_decomp: ∀n,l,qin,cin,qout,cout,mv. + match_in_table (S n) qin cin qout cout mv l → + ∃l1,l2. l = l1@(mk_tuple qin cin qout cout mv)@l2 ∧ + (∃q.|l1| = (tuple_length (S n))*q) ∧ + tuple_TM (S n) (mk_tuple qin cin qout cout mv). -if current (* x *) = # - then - else if x = 0 - then move_right; ---- - adv_to_mark_r; - if current (* x0 *) = 0 - then advance_mark ---- - adv_to_mark_l; - advance_mark - else STOP - else x = 1 (* analogo *) +lemma match_to_tuples_list: ∀n,h,l,qin,cin,qout,cout,mv. + match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_list n h l)) → + ∃p. p = mk_tuple qin cin qout cout mv ∧ mem ? p (tuples_list n h l). +#n #h #l #qin #cin #qout #cout #mv #Hmatch +@(ex_intro … (mk_tuple qin cin qout cout mv)) % // +cases (match_decomp … Hmatch) #l1 * #l2 * * #Hflat #Hlen #Htuple +@(flatten_to_mem … Hflat … Hlen) + [// + |#x #memx @length_of_tuple + cases (mem_map ????? memx) #t * #memt #Ht Hfalse in Hc; #Htf destruct (Htf) - | * #_ #Hta cases (tech_split STape (λc.is_bar (\fst c)) rs1) - [ #H1 lapply (Hta rs1 〈grid,false〉 rs2 (refl ??) ? ?) - [ * #x #b #Hx whd in ⊢ (??%?); >(Hrs1' … Hx) >(H1 … Hx) % - | % - | -Hta #Hta cases Hright - [ * #tb * whd in ⊢ (%→?); #Hcurrent - @False_ind cases (Hcurrent 〈grid,false〉 ?) - [ normalize #Hfalse destruct (Hfalse) - | >Hta % ] - | * #tb * whd in ⊢ (%→?); #Hcurrent - cases (Hcurrent 〈grid,false〉 ?) - [ #_ #Htb whd in ⊢ (%→?); #Houtc - %2 % - [ @H1 - | >Houtc >Htb >Hta % ] - | >Hta % ] - ] - ] - | * #rs3 * #c0 * #rs4 * * #Hc0 #Hsplit #Hrs3 - % @(ex_intro ?? rs3) @(ex_intro ?? rs4) - lapply (Hta rs3 c0 (rs4@〈grid,false〉::rs2) ???) - [ #x #Hrs3' whd in ⊢ (??%?); >Hsplit in Hrs1;>Hsplit in Hrs3; - #Hrs3 #Hrs1 >(Hrs1 …) [| @memb_append_l1 @Hrs3'|] - >(Hrs3 … Hrs3') @Hrs1' >Hsplit @memb_append_l1 // - | whd in ⊢ (??%?); >Hc0 % - | >Hsplit >associative_append % ] -Hta #Hta - cases Hright - [ * #tb * whd in ⊢ (%→?); #Hta' - whd in ⊢ (%→?); #Htb - cases (Hta' c0 ?) - [ #_ #Htb' >Htb' in Htb; #Htb - generalize in match Hsplit; -Hsplit - cases rs4 in Hta; - [ #Hta #Hsplit >(Htb … Hta) - >(?:c0 = 〈bar,false〉) - [ @(ex_intro ?? grid) @(ex_intro ?? false) - % [ % [ % - [(* Hsplit *) @daemon |(*Hrs3*) @daemon ] | % ] | % ] - | (* Hc0 *) @daemon ] - | #r5 #rs5 >(eq_pair_fst_snd … r5) - #Hta #Hsplit >(Htb … Hta) - >(?:c0 = 〈bar,false〉) - [ @(ex_intro ?? (\fst r5)) @(ex_intro ?? (\snd r5)) - % [ % [ % [ (* Hc0, Hsplit *) @daemon | (*Hrs3*) @daemon ] | % ] - | % ] | (* Hc0 *) @daemon ] ] | >Hta % ] - | * #tb * whd in ⊢ (%→?); #Hta' - whd in ⊢ (%→?); #Htb - cases (Hta' c0 ?) - [ #Hfalse @False_ind >Hfalse in Hc0; - #Hc0 destruct (Hc0) - | >Hta % ] -]]]] +axiom append_eq_tech1 : + ∀A,l1,l2,l3,l4.l1@l2 = l3@l4 → |l1| < |l3| → ∃la:list A.l1@la = l3. +axiom append_eq_tech2 : + ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false → ∃la:list A.l3 = l1@a::la. +axiom list_decompose_cases : + ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → ∃la,lb:list A.l3 = la@a::lb ∨ l4 = la@a::lb. +axiom list_decompose_l : + ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false → + ∃la,lb.l2 = la@lb ∧ l3 = l1@a::la.*) + +lemma list_decompose_r : + ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l3 = false → + ∃la,lb.l1 = la@lb ∧ l4 = lb@a::l2. +#A #l1 #l2 #l3 generalize in match l1; generalize in match l2; elim l3 + [normalize #l1 #l2 #l4 #a #H #_ @(ex_intro … []) @(ex_intro … l2) /2/ + |#b #tl #Hind #l1 #l2 #l4 #a cases l2 + [normalize #Heq destruct >(\b (refl … b)) normalize #Hfalse destruct + |#c #tl2 whd in ⊢ ((??%%)→?); #Heq destruct #Hmema + cases (Hind l1 tl2 l4 a ??) + [#l5 * #l6 * #eql #eql4 + @(ex_intro … (b::l5)) @(ex_intro … l6) % /2/ + |@e0 + |cases (true_or_false (memb ? a tl)) [2://] + #H @False_ind @(absurd ?? not_eq_true_false) + Ht + normalize #Hfalse destruct (Hfalse) +| #t0 #T0 #Htuple0 #Htable0 #_ #Heq + lapply (append_l2_injective ?????? Heq) + [ >(length_of_tuple … Htuple) >(length_of_tuple … Htuple0) % ] + -Heq #Heq destruct (Heq) // ] qed. -definition init_current ≝ - seq ? (adv_to_mark_l ? (is_marked ?)) - (seq ? (clear_mark ?) - (seq ? (adv_to_mark_l ? (λc:STape.is_grid (\fst c))) - (seq ? (move_r ?) (mark ?)))). - -definition R_init_current ≝ λt1,t2. - ∀l1,c,l2,b,l3,c1,rs,c0,b0. no_marks l1 → no_grids l2 → is_grid c = false → - Some ? 〈c0,b0〉 = option_hd ? (reverse ? (〈c,true〉::l2)) → - t1 = midtape STape (l1@〈c,true〉::l2@〈grid,b〉::l3) 〈c1,false〉 rs → - t2 = midtape STape (〈grid,b〉::l3) 〈c0,true〉 - ((tail ? (reverse ? (l1@〈c,false〉::l2))@〈c1,false〉::rs)). - -lemma sem_init_current : Realize ? init_current R_init_current. -#intape -cases (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?)) - (sem_seq ????? (sem_clear_mark ?) - (sem_seq ????? (sem_adv_to_mark_l ? (λc:STape.is_grid (\fst c))) - (sem_seq ????? (sem_move_r ?) (sem_mark ?)))) intape) -#k * #outc * #Hloop #HR -@(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -cases HR -HR #ta * whd in ⊢ (%→?); #Hta -* #tb * whd in ⊢ (%→?); #Htb -* #tc * whd in ⊢ (%→?); #Htc -* #td * whd in ⊢ (%→%→?); #Htd #Houtc -#l1 #c #l2 #b #l3 #c1 #rs #c0 #b0 #Hl1 #Hl2 #Hc #Hc0 #Hintape -cases (Hta … Hintape) [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] --Hta * #_ #Hta lapply (Hta l1 〈c,true〉 ? (refl ??) ??) [@Hl1|%] --Hta #Hta lapply (Htb … Hta) -Htb #Htb cases (Htc … Htb) [ >Hc -Hc * #Hc destruct (Hc) ] --Htc * #_ #Htc lapply (Htc … (refl ??) (refl ??) ?) [@Hl2] --Htc #Htc lapply (Htd … Htc) -Htd ->reverse_append >reverse_cons ->reverse_cons in Hc0; cases (reverse … l2) -[ normalize in ⊢ (%→?); #Hc0 destruct (Hc0) - #Htd >(Houtc … Htd) % -| * #c2 #b2 #tl2 normalize in ⊢ (%→?); - #Hc0 #Htd >(Houtc … Htd) - whd in ⊢ (???%); destruct (Hc0) - >associative_append >associative_append % +lemma match_in_table_to_tuple : + ∀n,T,qin,cin,qout,cout,mv. + match_in_table n qin cin qout cout mv T → table_TM n T → + tuple_TM n (mk_tuple qin cin qout cout mv). +#n #T #qin #cin #qout #cout #mv #Hmatch elim Hmatch +[ // +| #qin0 #cin0 #qout0 #cout0 #mv0 #tb #Htuple #Hmatch #IH #Htable + @IH @(table_invert_r ???? Htable) @Htuple ] qed. -definition match_tuple_step ≝ - ifTM ? (test_char ? (λc:STape.¬ is_grid (\fst c))) - (single_finalTM ? - (seq ? compare - (ifTM ? (test_char ? (λc:STape.is_grid (\fst c))) - (nop ?) - (seq ? mark_next_tuple - (ifTM ? (test_char ? (λc:STape.is_grid (\fst c))) - (mark ?) (seq ? (move_l ?) init_current) tc_true)) tc_true))) - (nop ?) tc_true. - -definition R_match_tuple_step_true ≝ λt1,t2. - ∀ls,c,l1,l2,c1,l3,l4,rs,n. - is_bit c = true → only_bits l1 → no_grids l2 → is_bit c1 = true → - only_bits l3 → n = |l2| → |l2| = |l3| → - table_TM (S n) (〈c1,true〉::l3@〈comma,false〉::l4) → - t1 = midtape STape (〈grid,false〉::ls) 〈c,true〉 - (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs) → - (* facciamo match *) - (〈c,true〉::l2 = 〈c1,true〉::l3 ∧ - t2 = midtape ? (reverse ? l2@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉 - (l2@〈bar,false〉::〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs)) - ∨ - (* non facciamo match e marchiamo la prossima tupla *) - (〈c,true〉::l2 ≠ 〈c1,true〉::l3 ∧ - ∃c2,l5,l6,l7.l4 = l5@〈bar,false〉::〈c2,false〉::l6@〈comma,false〉::l7 ∧ - (* condizioni su l5 l6 l7 *) - t2 = midtape STape (〈grid,false〉::ls) 〈c,true〉 - (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,true〉::l3@〈comma,false〉:: - l5@〈bar,false〉::〈c2,true〉::l6@〈comma,false〉::l7)) - ∨ - (* non facciamo match e non c'è una prossima tupla: - non specifichiamo condizioni sul nastro di output, perché - non eseguiremo altre operazioni, quindi il suo formato non ci interessa *) - (〈c,true〉::l2 ≠ 〈c1,true〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉). - -definition R_match_tuple_step_false ≝ λt1,t2. - ∀ls,c,rs.t1 = midtape STape ls c rs → is_grid (\fst c) = true ∧ t2 = t1. \ No newline at end of file +lemma match_in_table_append : + ∀n,T,qin,cin,qout,cout,mv,t. + tuple_TM n t → + match_in_table n qin cin qout cout mv (t@T) → + t = mk_tuple qin cin qout cout mv ∨ match_in_table n qin cin qout cout mv T. +#n #T #qin #cin #qout #cout #mv #t #Ht #Hmatch inversion Hmatch +[ #T0 #H #H1 % >(append_l1_injective … H1) // + >(length_of_tuple … Ht) >(length_of_tuple … H) % +| #qin0 #cin0 #qout0 #cout0 #mv0 #T0 #H #H1 #_ #H2 %2 + >(append_l2_injective … H2) // >(length_of_tuple … Ht) >(length_of_tuple … H) % +] +qed. + +lemma generic_match_to_match_in_table_tech : + ∀n,t,T0,T1,T2.tuple_TM n t → table_TM n (T1@〈bar,false〉::T2) → + t@T0 = T1@〈bar,false〉::T2 → T1 = [] ∨ ∃T3.T1 = t@T3. +#n #t #T0 #T1 #T2 #Ht cases T1 +[ #_ #_ % % +| normalize #c #T1c #Htable #Heq %2 + cases Ht in Heq; #qin * #cin * #qout * #cout * #mv ********** + #Hqin1 #Hqout1 #Hqin2 #Hqout2 #Hcin #Hcout #Hmv #Hcoutmv #Hqinlen #Hqoutlen + #Heqt >Heqt whd in ⊢ (??%%→?); #Ht lapply (cons_injective_r ????? Ht) + #Ht' cases (list_decompose_r STape … (sym_eq … Ht') ?) + [ #la * #lb * #HT1c #HT0 %{lb} >HT1c @(eq_f2 ??? (append ?) (c::la)) // + >HT0 in Ht'; >HT1c >associative_append in ⊢ (???%→?); #Ht' + <(append_l1_injective_r … Ht') // <(cons_injective_l ????? Ht) % + |@(noteq_to_eqnot ? true) @(not_to_not … not_eq_true_false) #Hbar @sym_eq + cases (memb_append … Hbar) -Hbar #Hbar + [@(Hqin2 … Hbar) + |cases (orb_true_l … Hbar) -Hbar + [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcin + |whd in ⊢ ((??%?)→?); #Hbar cases (memb_append … Hbar) -Hbar #Hbar + [@(Hqout2 … Hbar) + |cases (orb_true_l … Hbar) -Hbar + [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcout + |#Hbar cases (orb_true_l … Hbar) -Hbar + [whd in ⊢ ((??%?)→?); #Hbar @Hbar + |#Hbar lapply (memb_single … Hbar) -Hbar #Hbar destruct (Hbar) @Hmv + ] + ] + ] + ] + ] + ] +qed. + +lemma generic_match_to_match_in_table : + ∀n,T.table_TM n T → + ∀qin,cin,qout,cout,mv.|qin| = n → |qout| = n → + only_bits qin → only_bits qout → + bit_or_null (\fst cin) = true → bit_or_null (\fst cout) = true → + bit_or_null (\fst mv) = true → + ∀t1,t2. + T = (t1@〈bar,false〉::qin@cin::〈comma,false〉::qout@cout::〈comma,false〉::[mv])@t2 → + match_in_table n qin cin qout cout mv T. +#n #T #Htable #qin #cin #qout #cout #mv #Hlenqin #Hlenqout +#Hqinbits #Hqoutbits #Hcin #Hcout #Hmv +elim Htable +[ * [ #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) + | #c0 #t1 #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] +| #tuple #T0 #H1 #Htable0#IH #t1 #t2 #HT cases H1 #qin0 * #cin0 * #qout0 * #cout0 * #mv0 + * * * * * * * * * * + #Hqin0marks #Hqout0marks #Hqin0bits #Hqout0bits #Hcin0 #Hcout0 #Hmv0 #Hcout0mv0 + #Hlenqin0 #Hlenqout0 #Htuple + lapply (generic_match_to_match_in_table_tech n ? T0 t1 + (qin@cin::〈comma,false〉::qout@[cout;〈comma,false〉;mv]@t2) H1) #Htmp + >Htuple in H1; #H1 + lapply (ttm_cons … T0 H1 Htable0) HT + >associative_append normalize >associative_append normalize + >associative_append #Htable cases (Htmp Htable ?) + [ #Ht1 >Htuple in HT; >Ht1 normalize in ⊢ (??%%→?); + >associative_append >associative_append #HT + cut (qin0 = qin ∧ (〈cin0,false〉 = cin ∧ (qout0 = qout ∧ + (〈cout0,false〉 = cout ∧ (〈mv0,false〉 = mv ∧ T0 = t2))))) + [ lapply (cons_injective_r ????? HT) -HT #HT + lapply (append_l1_injective … HT) [ >Hlenqin @Hlenqin0 ] + #Hqin % [ @Hqin ] -Hqin + lapply (append_l2_injective … HT) [ >Hlenqin @Hlenqin0 ] -HT #HT + lapply (cons_injective_l ????? HT) #Hcin % [ @Hcin ] -Hcin + lapply (cons_injective_r ????? HT) -HT #HT + lapply (cons_injective_r ????? HT) -HT + >associative_append >associative_append #HT + lapply (append_l1_injective … HT) [ >Hlenqout @Hlenqout0 ] + #Hqout % [ @Hqout ] -Hqout + lapply (append_l2_injective … HT) [ >Hlenqout @Hlenqout0 ] -HT normalize #HT + lapply (cons_injective_l ????? HT) #Hcout % [ @Hcout ] -Hcout + lapply (cons_injective_r ????? HT) -HT #HT + lapply (cons_injective_r ????? HT) -HT #HT + lapply (cons_injective_l ????? HT) #Hmv % [ @Hmv ] -Hmv + @(cons_injective_r ????? HT) ] + -HT * #Hqin * #Hcin * #Hqout * #Hcout * #Hmv #HT0 + >(?:〈bar,false〉::qin0@(〈cin0,false〉::〈comma,false〉::qout0@ + [〈cout0,false〉;〈comma,false〉;〈mv0,false〉])@T0 = tuple@T0) + [ >Htuple >Hqin >Hqout >Hcin >Hcout >Hmv % // + | >Htuple normalize >associative_append normalize >associative_append + normalize >associative_append % ] + | * #T3 #HT3 >HT3 in HT; >associative_append; >associative_append #HT + lapply (append_l2_injective … HT) // -HT #HT %2 // + @(IH T3 t2) >HT >associative_append % + |>HT >associative_append normalize >associative_append normalize + >associative_append % ] +] +qed.