From 7f06a5a1b90a9fb9f1742cb8a469821148e0b9f9 Mon Sep 17 00:00:00 2001 From: Wilmer Ricciotti Date: Tue, 15 May 2012 17:01:42 +0000 Subject: [PATCH] Progress --- matita/matita/lib/turing/universal/tuples.ma | 62 +++++++++++++++----- 1 file changed, 47 insertions(+), 15 deletions(-) diff --git a/matita/matita/lib/turing/universal/tuples.ma b/matita/matita/lib/turing/universal/tuples.ma index 83e7337a0..984989b11 100644 --- a/matita/matita/lib/turing/universal/tuples.ma +++ b/matita/matita/lib/turing/universal/tuples.ma @@ -311,6 +311,12 @@ lemma some_option_hd: ∀A.∀l:list A.∀a.∃b. #A #l #a cases l normalize /2/ qed. +axiom tech_split2 : ∀A,l1,l2,l3,l4,x. + memb A x l1 = false → memb ? x l3 = false → + l1@x::l2 = l3@x::l4 → l1 = l3 ∧ l2 = l4. + +axiom injective_append : ∀A,l.injective … (λx.append A x l). + lemma sem_match_tuple_step: accRealize ? match_tuple_step (inr … (inl … (inr … 0))) R_match_tuple_step_true R_match_tuple_step_false. @@ -423,28 +429,54 @@ lemma sem_match_tuple_step: >reverse_reverse #Htapeout % [@Hnoteq] @(ex_intro … d2) - cut (∃rs31,rs32.rs3 = rs31@〈comma,false〉::rs32) - [@daemon] * #rs31 * #rs32 #Hrs3 + cut (∃rs32.rs3 = lc@〈comma,false〉::rs32) + [ (*cases (tech_split STape (λc.c == 〈bar,false〉) l4) + [ + | * #l41 * * #cbar #bfalse * #l42 * * #Hbar #Hl4 #Hl41 + @(ex_intro ?? l41) >Hl4 in Heq1; #Heq1 + + cut (sublist … lc l3) + [ #x #Hx cases la in H3; + [ normalize #H3 destruct (H3) @Hx + | #p #la' normalize #Hla' destruct (Hla') + @memb_append_l2 @memb_cons @Hx ] ] #Hsublist*) + @daemon] + * #rs32 #Hrs3 (* cut (〈c1,false〉::l3@〈comma,false〉::l4= la@〈d',false〉::rs3@〈bar,false〉::〈d2,b2〉::rs3') [@daemon] #Hcut *) + cut (l4=rs32@〈bar,false〉::〈d2,b2〉::rs3') + [ >Hrs3 in Heq1; @daemon ] #Hl4 @(ex_intro … rs32) @(ex_intro … rs3') % - [>Hrs3 in Heq1; @daemon + [(* by showing b2 = false: @Hl4 *) @daemon |>Htapeout @eq_f2 [@daemon - |(associative_append …la)

associative_append >associative_append + |(*>Hrs3 *)>append_cons + > (?:l1@〈grid,false〉::l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs + = (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::rs32@[〈bar,false〉])@〈d2,true〉::rs3'@〈grid,false〉::rs) + [|>associative_append normalize + >associative_append normalize + >associative_append normalize + >associative_append normalize + % ] + @eq_f2 [|%] + @(injective_append … (〈d2,b2〉::rs3')) + >(?:(la@[〈c',false〉])@((((lb@[〈grid,false〉])@l2@[〈bar,false〉])@la)@[〈d',false〉])@rs3 + =((la@〈c',false〉::lb)@([〈grid,false〉]@l2@[〈bar,false〉]@la@[〈d',false〉]@rs3))) + [|>associative_append >associative_append + >associative_append >associative_append >associative_append + >associative_append >associative_append % ] +

associative_append >associative_append >associative_append - >associative_append normalize @eq_f @eq_f @eq_f @eq_f - cut (true = b2) [@daemon] #Hcut >Hcut - cut (∀A,l1,l2,l3. l1 = l2 → l1@l3 = - (associative_append …la) - cut (true = b2 - - + @eq_f normalize @eq_f >associative_append + >associative_append @eq_f normalize @eq_f + >(append_cons ? 〈d',false〉) >associative_append + Hl4 associative_append normalize + >associative_append normalize % + ] + ] ] ] ] -- 2.39.2