From 2eef5f7f15de5fd3820075470c2937dba2012da6 Mon Sep 17 00:00:00 2001 From: Wilmer Ricciotti Date: Fri, 1 Jun 2012 14:40:49 +0000 Subject: [PATCH] Finalized copy sub-machine of the universal turing machine. Some new results on lists. --- matita/matita/lib/basics/lists/list.ma | 73 +++++- matita/matita/lib/turing/universal/copy.ma | 249 ++++++++++++++++--- matita/matita/lib/turing/universal/marks.ma | 30 +-- matita/matita/lib/turing/universal/tuples.ma | 13 +- 4 files changed, 293 insertions(+), 72 deletions(-) diff --git a/matita/matita/lib/basics/lists/list.ma b/matita/matita/lib/basics/lists/list.ma index edc049738..a20a89108 100644 --- a/matita/matita/lib/basics/lists/list.ma +++ b/matita/matita/lib/basics/lists/list.ma @@ -84,6 +84,14 @@ theorem nil_to_nil: ∀A.∀l1,l2:list A. #A #l1 #l2 #isnil @(nil_append_elim A l1 l2) /2/ qed. +lemma cons_injective_l : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → a1 = a2. +#A #a1 #a2 #l1 #l2 #Heq destruct // +qed. + +lemma cons_injective_r : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → l1 = l2. +#A #a1 #a2 #l1 #l2 #Heq destruct // +qed. + (**************************** iterators ******************************) let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝ @@ -183,16 +191,73 @@ lemma length_map: ∀A,B,l.∀f:A→B. length ? (map ?? f l) = length ? l. #A #B #l #f elim l // #a #tl #Hind normalize // qed. +lemma length_reverse: ∀A.∀l:list A. + |reverse A l| = |l|. +#A #l elim l // #a #l0 #IH >reverse_cons >length_append normalize // +qed. + +(****************** traversing two lists in parallel *****************) +lemma list_ind2 : + ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop. + length ? l1 = length ? l2 → + (P [] []) → + (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → + P l1 l2. +#T1 #T2 #l1 #l2 #P #Hl #Pnil #Pcons +generalize in match Hl; generalize in match l2; +elim l1 +[#l2 cases l2 // normalize #t2 #tl2 #H destruct +|#t1 #tl1 #IH #l2 cases l2 + [normalize #H destruct + |#t2 #tl2 #H @Pcons @IH normalize in H; destruct // ] +] +qed. + +lemma list_cases2 : + ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop. + length ? l1 = length ? l2 → + (l1 = [] → l2 = [] → P) → + (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P. +#T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl) +[ #Pnil #Pcons @Pnil // +| #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ] +qed. + +(*********************** properties of append ***********************) +lemma append_l1_injective : + ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l1 = l2. +#a #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen) // +#tl1 #tl2 #hd1 #hd2 #IH normalize #Heq destruct @eq_f /2/ +qed. + +lemma append_l2_injective : + ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l3 = l4. +#a #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen) normalize // +#tl1 #tl2 #hd1 #hd2 #IH normalize #Heq destruct /2/ +qed. + +lemma append_l1_injective_r : + ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l1 = l2. +#a #l1 #l2 #l3 #l4 #Hlen #Heq lapply (eq_f … (reverse ?) … Heq) +>reverse_append >reverse_append #Heq1 +lapply (append_l2_injective … Heq1) [ // ] #Heq2 +lapply (eq_f … (reverse ?) … Heq2) // +qed. + +lemma append_l2_injective_r : + ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l3 = l4. +#a #l1 #l2 #l3 #l4 #Hlen #Heq lapply (eq_f … (reverse ?) … Heq) +>reverse_append >reverse_append #Heq1 +lapply (append_l1_injective … Heq1) [ // ] #Heq2 +lapply (eq_f … (reverse ?) … Heq2) // +qed. + lemma length_rev_append: ∀A.∀l,acc:list A. |rev_append ? l acc| = |l|+|acc|. #A #l elim l // #a #tl #Hind normalize #acc >Hind normalize // qed. -lemma length_reverse: ∀A.∀l:list A. - |reverse A l| = |l|. -// qed. - (****************************** mem ********************************) let rec mem A (a:A) (l:list A) on l ≝ match l with diff --git a/matita/matita/lib/turing/universal/copy.ma b/matita/matita/lib/turing/universal/copy.ma index ae6f0e6cd..c6d7e7911 100644 --- a/matita/matita/lib/turing/universal/copy.ma +++ b/matita/matita/lib/turing/universal/copy.ma @@ -50,9 +50,49 @@ definition R_copy_step_subcase ≝ (x = c ∧ t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x,false〉::l2) 〈a,true〉 l3) ∨ (x ≠ c ∧ RelseM t1 t2). -axiom sem_copy_step_subcase : - ∀alpha,c,elseM,RelseM. +lemma sem_copy_step_subcase : + ∀alpha,c,elseM,RelseM. Realize ? elseM RelseM → Realize ? (copy_step_subcase alpha c elseM) (R_copy_step_subcase alpha c RelseM). +#alpha #c #elseM #RelseM #HelseM #intape +cases (sem_if ? (test_char ? (λx. x == 〈c,true〉)) ?????? tc_true (sem_test_char ? (λx.x == 〈c,true〉)) + (sem_seq ????? (sem_adv_mark_r alpha) + (sem_seq ????? (sem_move_l …) + (sem_seq ????? (sem_adv_to_mark_l … (is_marked alpha)) + (sem_seq ????? (sem_write ? 〈c,false〉) + (sem_seq ????? (sem_move_r …) + (sem_seq ????? (sem_mark …) + (sem_seq ????? (sem_move_r …) (sem_adv_to_mark_r … (is_marked alpha))))))))) + HelseM intape) +#k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop +#a #l1 #x0 #a0 #l2 #x #l3 #Hintape #Hl1marks cases HR -HR +[ * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta … (refl ??)) -Hta #Hx #Hta + * #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) -Hta -Htb #Htb + * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htb -Htc #Htc + * #td * whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd + [ >Htc * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] + * #_ #Htd lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ? (refl ??) ?) -Htd + [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1 [@(Hl1marks ? Hx1)|>(memb_single … Hx1) %] + | normalize >associative_append % ] #Htd + * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd -Htc #Hte + * #tf * whd in ⊢ (%→?); #Htf lapply (Htf … Hte) -Hte -Htf >reverse_append #Htf + * #tg * whd in ⊢ (%→?); #Htg lapply (Htg … Htf) -Htf -Htg >reverse_single #Htg + * #th * whd in ⊢ (%→?); #Hth lapply (Hth … Htg) -Htg -Hth + generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1) + [ #Hl1marks #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc + [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] + * #_ #Houtc lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc + #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] @Houtc + | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single + #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc + [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ] + * #_ #Houtc lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc + [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1 [ @Hl1marks @memb_append_l1 @daemon | >(memb_single … Hx1) % ] + | normalize >associative_append % ] + #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc + >reverse_append >reverse_reverse >associative_append >associative_append % ] +| * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta ? (refl ??)) -Hta + #Hxc #Hta >Hta #Houtc %2 % // lapply (\Pf Hxc) @not_to_not #Heq >Heq % ] +qed. (* if current = 0,tt @@ -101,27 +141,49 @@ definition R_nocopy_subcase ≝ t2 = midtape ? (〈x,false〉::l1@〈a0,true〉::〈x0,false〉::l2) 〈a,true〉 l3) ∨ (x ≠ null ∧ t2 = t1). -axiom sem_nocopy_subcase : Realize ? nocopy_subcase R_nocopy_subcase. -(* #intape +lemma sem_nocopy_subcase : Realize ? nocopy_subcase R_nocopy_subcase. +#intape cases (sem_if ? (test_char ? (λx:STape.x == 〈null,true〉)) ?????? tc_true - (sem_test_char ? (λx:STape.x == 〈null,true〉)) - (sem_seq … (sem_adv_mark_r …) - (sem_seq … (sem_move_l …) - (sem_seq … (sem_adv_to_mark_l … (is_marked ?)) - (sem_seq … (sem_adv_mark_r …) - (sem_seq … (sem_move_r …) (sem_adv_to_mark_r … (is_marked ?)) - ))))) (sem_nop ?) intape) + (sem_test_char ? (λx:STape.x == 〈null,true〉)) + (sem_seq … (sem_adv_mark_r …) + (sem_seq … (sem_move_l …) + (sem_seq … (sem_adv_to_mark_l … (is_marked ?)) + (sem_seq … (sem_adv_mark_r …) + (sem_seq … (sem_move_r …) + (sem_adv_to_mark_r … (is_marked ?))))))) (sem_nop ?) intape) #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [@Hloop] -Hloop -cases HR -HR -[| * #ta * whd in ⊢ (%→%→?); #Hta #Houtc - #ls #x #rs #Hintape %2 >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta #Hx #Hta % - [ lapply (\Pf Hx) @not_to_not #Hx' >Hx' % - | Hintape #Hta cases (Hta … (refl ??)) -Hta #Hx #Hta + * #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) -Hta -Htb #Htb + * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htb -Htc #Htc + * #td * whd in ⊢ (%→?); #Htd cases (Htd … Htc) -Htd + [ >Htc * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] + * #_ #Htd lapply (Htd (l1@[〈a0,false〉]) 〈x0,true〉 l2 ? (refl ??) ?) -Htd + [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1 [@(Hl1marks ? Hx1)|>(memb_single … Hx1) %] + | normalize >associative_append % ] >reverse_append #Htd + * #te * whd in ⊢ (%→?); #Hte lapply (Hte … Htd) -Hte -Htd -Htc #Hte + * #tf * whd in ⊢ (%→?); #Htf lapply (Htf … Hte) -Hte -Htf + generalize in match Hl1marks; -Hl1marks @(list_elim_left … l1) + [ #Hl1marks #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc + [ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] + * #_ #Houtc lapply (Houtc [] ?? (refl ??) (refl ??) Hl1marks) -Houtc + #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] @Houtc + | -l1 #c1 #l1 #_ #Hl1marks >reverse_append >reverse_single + #Hth whd in ⊢ (%→?); #Houtc cases (Houtc … Hth) -Houtc + [ * >Hl1marks [ #Hfalse destruct (Hfalse) ] @memb_append_l2 @memb_hd ] + * #_ #Houtc lapply (Houtc (reverse ? l1@[〈x,false〉]) 〈a,true〉 l3 ? (refl ??) ?) -Houtc + [ #x1 #Hx1 cases (memb_append … Hx1) -Hx1 #Hx1 [ @Hl1marks @memb_append_l1 @daemon | >(memb_single … Hx1) % ] + | normalize >associative_append % ] + #Houtc lapply (\P Hx) -Hx #Hx destruct (Hx) % % [%] >Houtc + >reverse_append >reverse_reverse >associative_append >associative_append % ] +| * #ta * whd in ⊢ (%→?); >Hintape #Hta cases (Hta ? (refl ??)) -Hta + #Hxc #Hta >Hta whd in ⊢ (%→?); #Houtc %2 % + [ lapply (\Pf Hxc) @not_to_not #Heq >Heq % + | @Houtc ] +qed. definition copy_step ≝ - ifTM ? (test_char STape (λc.is_bit (\fst c))) + ifTM ? (test_char STape (λc.bit_or_null (\fst c))) (single_finalTM ? (copy_step_subcase FSUnialpha (bit false) (copy_step_subcase FSUnialpha (bit true) nocopy_subcase))) (nop ?) @@ -144,9 +206,37 @@ definition R_copy_step_false ≝ λt1,t2. ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → bit_or_null (\fst c) = false ∧ t2 = t1. - -axiom sem_copy_step : + +lemma sem_copy_step : accRealize ? copy_step (inr … (inl … (inr … start_nop))) R_copy_step_true R_copy_step_false. +#intape +@(acc_sem_if_app … (sem_test_char ? (λc:STape.bit_or_null (\fst c))) … + (sem_copy_step_subcase FSUnialpha (bit false) … + (sem_copy_step_subcase FSUnialpha (bit true) … (sem_nocopy_subcase …))) + (sem_nop …)) +[ #t1 #t2 #t3 whd in ⊢ (%→%→?); #H1 #H2 #ls #c #rs #Ht1 >Ht1 in H1; #H1 + cases (H1 … (refl ??)) #Hc #Ht3 % [ @Hc ] + #a #l1 #x0 #a0 #l2 #l3 #Hls #Hrs #Hl1marks >Hls in Ht3; >Hrs #Ht3 + cases (H2 … Ht3 ?) + [ * #Hc' #Ht2 % %{false} % // Ht1 in H1; #H1 cases (H1 … (refl ??)) #_ #Ht3 cases (H1 ? (refl ??)) -H1 + #Hc #Ht3 % // +] +qed. (* 1) il primo carattere è marcato @@ -200,8 +290,6 @@ lemma inj_append_singleton_l2 : >reverse_append >reverse_append normalize #H1 destruct % qed. -axiom length_reverse : ∀A,l.|reverse A l| = |l|. - lemma wsem_copy0 : WRealize ? copy0 R_copy0. #intape #k #outc #Hloop lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop @@ -223,7 +311,7 @@ lapply (sem_while … sem_copy_step intape k outc Hloop) [%] -Hloop cases (Htc … Htb) -Htc #Hcbitnull #Htc % [ % #Hc' >Hc' in Hcbitnull; normalize #Hfalse destruct (Hfalse) ] cut (|l1| = |reverse ? l4|) [@daemon] #Hlen1 - @(list_cases_2 … Hlen1) + @(list_cases2 … Hlen1) [ (* case l1 = [] is discriminated because l1 contains at least comma *) #Hl1nil @False_ind >Hl1nil in Hl1; cases l1' normalize [ #Hl1 destruct normalize in Hcbitnull; destruct (Hcbitnull) @@ -410,19 +498,39 @@ lemma merge_cons : cases c2 /2/ qed. +lemma merge_bits : ∀l1,l2.|l1| = |l2| → only_bits l2 → merge_config l1 l2 = l2. +#l1 #l2 #Hlen @(list_ind2 … Hlen) // +#tl1 #tl2 #hd1 #hd2 #IH +>(eq_pair_fst_snd … hd1) >(eq_pair_fst_snd … hd2) #Hbits +change with (cons ???) in ⊢ (??%?); @eq_f2 +[ cases (\fst hd2) in Hbits; + [ #b #_ % + |*: #Hfalse lapply (Hfalse … (memb_hd …)) normalize #Hfalse1 destruct (Hfalse1) ] +| @IH #x #Hx @Hbits @memb_cons // ] +qed. + lemma merge_config_c_nil : ∀c.merge_config c [] = []. #c cases c normalize // qed. -axiom reverse_merge_config : +lemma reverse_merge_config : ∀c1,c2.|c1| = |c2| → reverse ? (merge_config c1 c2) = merge_config (reverse ? c1) (reverse ? c2). +#c1 #c2 <(length_reverse ? c1) <(length_reverse ? c2) #Hlen +<(reverse_reverse ? c1) in ⊢ (??%?); <(reverse_reverse ? c2) in ⊢ (??%?); +generalize in match Hlen; @(list_ind2 … Hlen) -Hlen // +#tl1 #tl2 #hd1 #hd2 #IH whd in ⊢ (??%%→?); #Hlen destruct (Hlen) -Hlen +<(length_reverse ? tl1) in e0; <(length_reverse ? tl2) #Hlen +>reverse_cons >reverse_cons >(merge_config_append ???? Hlen) +>reverse_append >(eq_pair_fst_snd ?? hd1) >(eq_pair_fst_snd ?? hd2) +whd in ⊢ (??%%); @eq_f2 // @IH // +qed. definition copy ≝ - seq STape (move_l …) (seq ? (adv_to_mark_l … (is_marked ?)) - (seq ? (clear_mark …) (seq ? (adv_to_mark_r … (is_marked ?)) (clear_mark …)))). + seq STape copy0 (seq ? (move_l …) (seq ? (adv_to_mark_l … (is_marked ?)) + (seq ? (clear_mark …) (seq ? (adv_to_mark_r … (is_marked ?)) (clear_mark …))))). (* s0, s1 = caratteri di testa dello stato @@ -438,10 +546,93 @@ definition R_copy ≝ λt1,t2. ∀ls,s0,s1,c0,c1,rs,l1,l3,l4. t1 = midtape STape (l3@〈grid,false〉::〈c0,false〉::l4@〈s0,true〉::ls) 〈s1,true〉 (l1@〈c1,false〉::〈comma,false〉::rs) → no_marks l1 → no_marks l3 → no_marks l4 → |l1| = |l4| → - only_bits (l4@[〈s0,true〉]) → only_bits (〈s1,true〉::l1) → + only_bits (l4@[〈s0,false〉]) → only_bits (〈s1,false〉::l1) → bit_or_null c0 = true → bit_or_null c1 = true → t2 = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l3@〈grid,false〉:: 〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls) 〈comma,false〉 rs. + +axiom sem_copy0 : Realize ? copy0 R_copy0. + +definition option_cons ≝ λA.λa:option A.λl. + match a with + [ None ⇒ l + | Some a' ⇒ a'::l ]. -axiom sem_copy : Realize ? copy R_copy. \ No newline at end of file +lemma sem_copy : Realize ? copy R_copy. +#intape +cases (sem_seq … (sem_copy0 …) + (sem_seq … (sem_move_l …) + (sem_seq … (sem_adv_to_mark_l … (is_marked ?)) + (sem_seq … (sem_clear_mark …) + (sem_seq … (sem_adv_to_mark_r … (is_marked ?)) (sem_clear_mark …))))) intape) +#k * #outc * #Hloop #HR %{k} %{outc} % [@Hloop] -Hloop +#ls #s0 #s1 #c0 #c1 #rs #l1 #l2 #l3 #Hintape #Hl1marks #Hl2marks #Hl3marks #Hlen +#Hbits1 #Hbits2 #Hc0bits #Hc1bits +cases HR -HR #ta * whd in ⊢ (%→?); #Hta +cut (ta = midtape STape (〈c1,false〉::reverse ? l1@〈s1,false〉::l2@〈grid,true〉:: + 〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls) + 〈comma,true〉 rs) +[lapply (Hta ls s1 s0 rs (l1@[〈c1,false〉;〈comma,false〉]) l2 (〈grid,false〉::〈c0,false〉::l3) ?) + [>associative_append in ⊢ (???(????%)); normalize in ⊢ (???(??%%%)); @Hintape ] + -Hta #Hta cases (Hta ??? (〈s1,false〉::l1@[〈c1,false〉]) false ? ? ?? (refl ??) ?) + [3: #x #Hx cases (memb_append … Hx) -Hx #Hx + [ @Hl1marks // + | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | >(memb_single … Hx) % ]] + |4: #x #Hx cases (memb_append … Hx) -Hx #Hx + [ @Hl2marks // + | cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) % | cases (orb_true_l … Hx) [-Hx #Hx >(\P Hx) % | @Hl3marks ] ] ] + |5: >length_append normalize >Hlen >commutative_plus % + |6: normalize >associative_append % + |7: #x #Hx cases (memb_append ?? (〈s1,false〉::l1) … Hx) -Hx #Hx + [ whd in ⊢ (??%?); >(Hbits2 … Hx) % + | >(memb_single … Hx) // ] + |8: #x #Hx cases (memb_append … Hx) -Hx #Hx + [ cases (orb_true_l … Hx) -Hx #Hx [ >(\P Hx) // | whd in ⊢ (??%?); >Hbits1 // @memb_append_l1 // ] + | >(memb_single … Hx) whd in ⊢ (??%?); >(Hbits1 〈s0,false〉) // @memb_append_l2 @memb_hd ] + | * #Hs1 @False_ind >Hs1 in Hbits2; #Hbits2 lapply (Hbits2 〈comma,false〉 ?) // + normalize #Hfalse destruct (Hfalse) + | * #Hs1 #Ht2 >Ht2 >reverse_cons >reverse_append >reverse_single @eq_f3 // + >merge_cons >merge_bits + [2: #x #Hx @Hbits2 cases (memb_append STape ? (reverse ? l1) ? Hx) -Hx #Hx + [@daemon | >(memb_single … Hx) @memb_hd ] + |3: >length_append >length_append >length_reverse >Hlen % ] + normalize >associative_append normalize >associative_append % + ] +] -Hta #Hta * #tb * whd in ⊢ (%→?); #Htb +lapply (Htb … Hta) -Htb #Htb change with (midtape ????) in Htb:(???%); +* #tc * whd in ⊢ (%→?); #Htc +cases (Htc … Htb) +[ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] +* #_ #Htc +lapply (Htc (reverse ? l1@〈s1,false〉::l2) 〈grid,true〉 + (〈merge_char c0 c1,false〉::reverse ? l1@〈s1,false〉::ls)???) +[ #x #Hx cases (memb_append … Hx) -Hx #Hx + [ @Hl1marks @daemon + | cases (orb_true_l … Hx) -Hx #Hx + [ >(\P Hx) % | @(Hl2marks … Hx) ] ] +| % +| whd in ⊢ (??%?); >associative_append % ] -Htc #Htc +* #td * whd in ⊢ (%→?); #Htd lapply (Htd … Htc) -Htd #Htd +* #te * whd in ⊢ (%→?); #Hte cases (Hte … Htd) -Hte -Htd +[ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] +* #_ #Hte +lapply (Hte (reverse ? (reverse ? l1@〈s1,false〉::l2)@[〈c1,false〉]) + 〈comma,true〉 rs ? (refl ??) ?) -Hte +[ >reverse_append >reverse_cons >reverse_reverse #x #Hx + cases (memb_append … Hx) -Hx #Hx + [ cases (memb_append … Hx) -Hx #Hx + [ cases (memb_append … Hx) -Hx #Hx + [ @daemon + | lapply (memb_single … Hx) -Hx #Hx >Hx % ] + | @(Hl1marks … Hx) ] + | lapply (memb_single … Hx) -Hx #Hx >Hx % ] +| >reverse_append >reverse_reverse >reverse_cons + >associative_append >associative_append >associative_append + >associative_append >associative_append % ] +#Hte whd in ⊢ (%→?); #Houtc lapply (Houtc … Hte) -Houtc #Houtc >Houtc +@eq_f3 // +>reverse_append >reverse_append >reverse_single >reverse_cons +>reverse_append >reverse_append >reverse_reverse >reverse_reverse +>reverse_single >associative_append >associative_append % +qed. \ No newline at end of file diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index 70000e399..46de2c4d6 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -864,32 +864,6 @@ definition R_compare := reverse ? la@ls) 〈d',false〉 (lc@〈comma,false〉::l2)). -lemma list_ind2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop. - length ? l1 = length ? l2 → - (P [] []) → - (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → - P l1 l2. -#T1 #T2 #l1 #l2 #P #Hl #Pnil #Pcons -generalize in match Hl; generalize in match l2; -elim l1 -[#l2 cases l2 // normalize #t2 #tl2 #H destruct -|#t1 #tl1 #IH #l2 cases l2 - [normalize #H destruct - |#t2 #tl2 #H @Pcons @IH normalize in H; destruct // ] -] -qed. - -lemma list_cases_2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop. - length ? l1 = length ? l2 → - (l1 = [] → l2 = [] → P) → - (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P. -#T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl) -[ #Pnil #Pcons @Pnil // -| #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ] -qed. - lemma wsem_compare : WRealize ? compare R_compare. #t #i #outc #Hloop lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] @@ -925,7 +899,7 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] |#b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1 #Heq destruct (Heq) #_ #Hrs cases Hleft -Hleft [2: * >Hc' #Hfalse @False_ind destruct ] * #_ - @(list_cases_2 … Hlen) + @(list_cases2 … Hlen) [ #Hbs #Hb0s generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?); -Hrs #Hrs normalize in Hrs; #Hleft cases (Hleft ????? Hrs ?) -Hleft [ * #Heqb #Htapeb cases (IH … Htapeb) -IH * #IH #_ #_ @@ -1002,5 +976,5 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] ] ]]]]] qed. - + axiom sem_compare : Realize ? compare R_compare. \ No newline at end of file diff --git a/matita/matita/lib/turing/universal/tuples.ma b/matita/matita/lib/turing/universal/tuples.ma index 9c6c7a09f..dc55b13a7 100644 --- a/matita/matita/lib/turing/universal/tuples.ma +++ b/matita/matita/lib/turing/universal/tuples.ma @@ -19,7 +19,8 @@ include "turing/universal/marks.ma". definition STape ≝ FinProd … FSUnialpha FinBool. definition only_bits ≝ λl. - ∀c.memb STape c l = true → is_bit (\fst c) = true. + ∀c.memb STape c l + = true → is_bit (\fst c) = true. definition only_bits_or_nulls ≝ λl. ∀c.memb STape c l = true → bit_or_null (\fst c) = true. @@ -81,16 +82,6 @@ inductive match_in_table (n:nat) (qin:list STape) (cin: STape) match_in_table n qin cin qout cout mv (mk_tuple qin0 cin0 qout0 cout0 mv0@tb). -axiom append_l1_injective : - ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l1 = l2. -axiom append_l2_injective : - ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l3 = l4. -axiom append_l1_injective_r : - ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l1 = l2. -axiom append_l2_injective_r : - ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l3 = l4. -axiom cons_injective_l : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → a1 = a2. -axiom cons_injective_r : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → l1 = l2. axiom tuple_len : ∀n,t.tuple_TM n t → |t| = 2*n+6. axiom append_eq_tech1 : ∀A,l1,l2,l3,l4.l1@l2 = l3@l4 → |l1| < |l3| → ∃la:list A.l1@la = l3. -- 2.39.2