X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Fmarks.ma;h=e82d577bc3ed666240af9ef6013d5f4b2f8d5f86;hb=d64a1790db147a15917f3c999dc5b35211dc5b56;hp=ddb99c96840378dca17ce6897a375c32264da318;hpb=01480f5f9d65c9f9800ea80b3cb7535c695d6a3f;p=helm.git diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index ddb99c968..e82d577bc 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -111,15 +111,31 @@ lemma sem_atmr_step : ] qed. +lemma dec_test: ∀alpha,rs,test. + decidable (∀c.memb alpha c rs = true → test c = false). +#alpha #rs #test elim rs + [%1 #n normalize #H destruct + |#a #tl cases (true_or_false (test a)) #Ha + [#_ %2 % #Hall @(absurd ?? not_eq_true_false) (\P eqca) @Ha |@Hall] + |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons // + ] + qed. + definition R_adv_to_mark_r ≝ λalpha,test,t1,t2. (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,rs. (t1 = midtape alpha ls c rs → ((test c = true ∧ t2 = t1) ∨ (test c = false ∧ - ∀rs1,b,rs2. rs = rs1@b::rs2 → + (∀rs1,b,rs2. rs = rs1@b::rs2 → test b = true → (∀x.memb ? x rs1 = true → test x = false) → - t2 = midtape ? (reverse ? rs1@c::ls) b rs2))). + t2 = midtape ? (reverse ? rs1@c::ls) b rs2) ∧ + ((∀x.memb ? x rs = true → test x = false) → + ∀a,l.reverse ? (c::rs) = a::l → + t2 = rightof alpha a (l@ls))))). definition adv_to_mark_r ≝ λalpha,test.whileTM alpha (atmr_step alpha test) atm2. @@ -149,20 +165,37 @@ lapply (sem_while … (sem_atmr_step alpha test) t i outc Hloop) [%] whd in ⊢((??%?)→?); #H destruct (H); |#ls #c #rs #Htapea %2 cases Hleft #ls0 * #a0 * #rs0 * * #Htapea' #Htest #Htapeb - >Htapea' in Htapea; #Htapea destruct (Htapea) % // * + >Htapea' in Htapea; #Htapea destruct (Htapea) % [ % // ] + [* [ #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #_ cases (proj2 ?? IH … Htapeb) [ * #_ #Houtc >Houtc >Htapeb % - | * #Hfalse >Hfalse in Htestb; #Htestb destruct (Htestb) ] + | * * >Htestb #Hfalse destruct (Hfalse) ] | #r1 #rs1 #b #rs2 #Hrs >Hrs in Htapeb; #Htapeb #Htestb #Hmemb cases (proj2 ?? IH … Htapeb) [ * #Hfalse >(Hmemb …) in Hfalse; [ #Hft destruct (Hft) | @memb_hd ] - | * #Htestr1 #H1 >reverse_cons >associative_append + | * * #Htestr1 #H1 #_ >reverse_cons >associative_append @H1 // #x #Hx @Hmemb @memb_cons // ] ] + |cases rs in Htapeb; normalize in ⊢ (%→?); + [#Htapeb #_ #a0 #l whd in ⊢ ((??%?)→?); #Hrev destruct (Hrev) + >Htapeb in IH; #IH cases (proj1 ?? IH … (refl …)) // + |#r1 #rs1 #Htapeb #Hmemb + cases (proj2 ?? IH … Htapeb) [ * >Hmemb [ #Hfalse destruct(Hfalse) ] @memb_hd ] + * #_ #H1 #a #l <(reverse_reverse … l) cases (reverse … l) + [#H cut (c::r1::rs1 = [a]) + [<(reverse_reverse … (c::r1::rs1)) >H //] + #Hrev destruct (Hrev) + |#a1 #l2 >reverse_cons >reverse_cons >reverse_cons + #Hrev cut ([c] = [a1]) + [@(append_l2_injective_r ?? (a::reverse … l2) … Hrev) //] + #Ha associative_append @H1 + [#x #membx @Hmemb @memb_cons @membx + |<(append_l1_injective_r ?? (a::reverse … l2) … Hrev) // + ] qed. lemma terminate_adv_to_mark_r : @@ -575,16 +608,17 @@ definition adv_both_marks ≝ λalpha. adv_to_mark_l (FinProd alpha FinBool) (is_marked alpha) · adv_mark_r alpha. -definition R_adv_both_marks ≝ - λalpha,t1,t2. - ∀l0,x,a,l1,x0. (∀c.memb ? c l1 = true → is_marked ? c = false) → - (∀l1',a0,l2. t1 = midtape (FinProd … alpha FinBool) - (l1@〈x,true〉::l0) 〈x0,true〉 (〈a0,false〉::l2) → - reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' → - t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)) ∧ - (t1 = midtape (FinProd … alpha FinBool) - (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 [ ] → - t2 = rightof ? 〈x0,false〉 (l1@〈a,false〉::〈x,true〉::l0)). +definition R_adv_both_marks ≝ λalpha,t1,t2. + ∀ls,x0,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs → + (rs = [ ] → (* first case: rs empty *) + t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧ + (∀l0,x,a,a0,b,l1,l1',l2. + ls = (l1@〈x,true〉::l0) → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = (〈a0,b〉::l2) → + reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' → + t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)). lemma sem_adv_both_marks : ∀alpha.Realize ? (adv_both_marks alpha) (R_adv_both_marks alpha). @@ -593,20 +627,22 @@ lemma sem_adv_both_marks : (sem_seq ????? (sem_move_l …) (sem_seq ????? (sem_adv_to_mark_l ? (is_marked ?)) (sem_adv_mark_r alpha))) …) -#intape #outtape * #tapea * #Hta * #tb * #Htb * #tc * #Htc #Hout -#l0 #x #a #l1 #x0 #Hmarks % - [#l1' #a0 #l2 #Hintape #Hrev @(proj1 ?? (proj1 ?? Hout … ) ? false) -Hout +#intape #outtape * #tapea * #Hta * #tb * #Htb * #tc * #Htc #Hout +#ls #c #rs #Hintape % + [#Hrs >Hrs in Hintape; #Hintape lapply (proj2 ?? (proj1 ?? Hta … ) … Hintape) -Hta #Hta + lapply (proj1 … Htb) >Hta -Htb #Htb lapply (Htb (refl …)) -Htb #Htb + lapply (proj1 ?? Htc) Htc @(proj2 ?? Hout …) Hrs in Hintape; >Hls #Hintape + @(proj1 ?? (proj1 ?? Hout … ) ? false) -Hout lapply (proj1 … (proj1 … Hta …) … Hintape) #Htapea lapply (proj2 … Htb … Htapea) -Htb whd in match (mk_tape ????) ; #Htapeb lapply (proj1 ?? (proj2 ?? (proj2 ?? Htc … Htapeb) (refl …))) -Htc #Htc change with ((?::?)@?) in match (cons ???); reverse_cons - >associative_append @Htc [%|%|@Hmarks] - |#Hintape lapply (proj2 ?? (proj1 ?? Hta … ) … Hintape) -Hta #Hta - lapply (proj1 … Htb) >Hta -Htb #Htb lapply (Htb (refl …)) -Htb #Htb - lapply (proj1 ?? Htc) Htc @(proj2 ?? Hout …) associative_append @Htc [%|%|@Hmarks] + ] qed. (* @@ -669,12 +705,21 @@ definition match_and_adv ≝ definition R_match_and_adv ≝ λalpha,f,t1,t2. - ∀l0,x,a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) → - t1 = midtape (FinProd … alpha FinBool) - (l1@〈a,false〉::〈x,true〉::l0) 〈x0,true〉 (〈a0,false〉::l2) → - (f 〈x0,true〉 = true ∧ t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (reverse ? l1@〈x0,false〉::〈a0,true〉::l2)) - ∨ (f 〈x0,true〉 = false ∧ - t2 = midtape ? (l1@〈a,false〉::〈x,true〉::l0) 〈x0,false〉 (〈a0,false〉::l2)). + ∀ls,x0,rs. + t1 = midtape (FinProd … alpha FinBool) ls 〈x0,true〉 rs → + ((* first case: (f 〈x0,true〉 = false) *) + (f 〈x0,true〉 = false) → + t2 = midtape (FinProd … alpha FinBool) ls 〈x0,false〉 rs) ∧ + ((f 〈x0,true〉 = true) → rs = [ ] → (* second case: rs empty *) + t2 = rightof (FinProd … alpha FinBool) 〈x0,false〉 ls) ∧ + ((f 〈x0,true〉 = true) → + ∀l0,x,a,a0,b,l1,l1',l2. + (* third case: we expect to have a mark on the left! *) + ls = (l1@〈x,true〉::l0) → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = 〈a0,b〉::l2 → + reverse ? (〈x0,false〉::l1) = 〈a,false〉::l1' → + t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1'@〈a0,true〉::l2)). lemma sem_match_and_adv : ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv alpha f). @@ -682,31 +727,52 @@ lemma sem_match_and_adv : cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape) #k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop -cases Hif -[ * #ta * whd in ⊢ (%→%→?); #Hta #Houtc - #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape >Hintape in Hta; - * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hf #Hta % % - [ @Hf | >append_cons >append_cons in Hta; #Hta @(proj1 ?? (Houtc …) …Hta) - [ #x #memx cases (memb_append …memx) - [@Hl1 | -memx #memx >(memb_single … memx) %] - |>reverse_cons >reverse_append % ] ] -| * #ta * whd in ⊢ (%→%→?); #Hta #Houtc - #l0 #x #a #l1 #x0 #a0 #l2 #Hl1 #Hintape >Hintape in Hta; - * #Hf #Hta %2 % [ @Hf % | >(proj2 ?? Houtc … Hta) % ] +(* +@(sem_if_app … (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?)) +#intape #outape #Htb * #H *) +cases Hif -Hif +[ * #ta * whd in ⊢ (%→%→?); * * #c * #Hcurrent #fc #Hta #Houtc + #ls #x #rs #Hintape >Hintape in Hcurrent; whd in ⊢ ((??%?)→?); #H destruct (H) % + [%[>fc #H destruct (H) + |#_ #Hrs >Hrs in Hintape; #Hintape >Hintape in Hta; #Hta + cases (Houtc … Hta) #Houtc #_ @Houtc // + ] + |#l0 #x0 #a #a0 #b #l1 #l1' #l2 #Hls #Hmarks #Hrs #Hrev >Hintape in Hta; #Hta + @(proj2 ?? (Houtc … Hta) … Hls Hmarks Hrs Hrev) + ] +| * #ta * * #H #Hta * #_ #Houtc #ls #c #rs #Hintape + >Hintape in H; #H lapply(H … (refl …)) #fc % + [%[#_ >Hintape in Hta; #Hta @(Houtc … Hta) + |>fc #H destruct (H) + ] + |>fc #H destruct (H) + ] ] qed. -(* - if x = c - then move_right; ---- - adv_to_mark_r; - if current (* x0 *) = 0 - then advance_mark ---- - adv_to_mark_l; - advance_mark - else STOP - else M -*) +definition R_match_and_adv_of ≝ + λalpha,t1,t2.current (FinProd … alpha FinBool) t1 = None ? → t2 = t1. + +lemma sem_match_and_adv_of : + ∀alpha,f.Realize ? (match_and_adv alpha f) (R_match_and_adv_of alpha). +#alpha #f #intape +cases (sem_if ? (test_char ? f) … tc_true (sem_test_char ? f) (sem_adv_both_marks alpha) (sem_clear_mark ?) intape) +#k * #outc * #Hloop #Hif @(ex_intro ?? k) @(ex_intro ?? outc) +% [ @Hloop ] -Hloop +cases Hif +[ * #ta * whd in ⊢ (%→%→?); #Hta #Houtc #Hcur + cases Hta * #x >Hcur * #Hfalse destruct (Hfalse) +| * #ta * whd in ⊢ (%→%→?); * #_ #Hta * #Houtc #_ (Houtc Hcur) % ] +qed. + +lemma sem_match_and_adv_full : + ∀alpha,f.Realize ? (match_and_adv alpha f) + (R_match_and_adv alpha f ∩ R_match_and_adv_of alpha). +#alpha #f #intape cases (sem_match_and_adv ? f intape) +#i * #outc * #Hloop #HR1 %{i} %{outc} % // % // +cases (sem_match_and_adv_of ? f intape) #i0 * #outc0 * #Hloop0 #HR2 +>(loop_eq … Hloop Hloop0) // +qed. definition comp_step_subcase ≝ λalpha,c,elseM. ifTM ? (test_char ? (λx.x == c)) @@ -715,17 +781,25 @@ definition comp_step_subcase ≝ λalpha,c,elseM. definition R_comp_step_subcase ≝ λalpha,c,RelseM,t1,t2. - ∀l0,x,rs.t1 = midtape (FinProd … alpha FinBool) l0 〈x,true〉 rs → + ∀ls,x,rs.t1 = midtape (FinProd … alpha FinBool) ls 〈x,true〉 rs → (〈x,true〉 = c → - ((∀c.memb ? c rs = true → is_marked ? c = false) → - ∀a,l. (a::l) = reverse ? (〈x,true〉::rs) → t2 = rightof (FinProd alpha FinBool) a (l@l0)) ∧ - ∀a,l1,x0,a0,l2. (∀c.memb ? c l1 = true → is_marked ? c = false) → - rs = 〈a,false〉::l1@〈x0,true〉::〈a0,false〉::l2 → - ((x = x0 → - t2 = midtape ? (〈x,false〉::l0) 〈a,true〉 (l1@〈x0,false〉::〈a0,true〉::l2)) ∨ - (x ≠ x0 → - t2 = midtape (FinProd … alpha FinBool) - (reverse ? l1@〈a,false〉::〈x,true〉::l0) 〈x0,false〉 (〈a0,false〉::l2)))) ∨ + ((* test true but no marks in rs *) + (∀c.memb ? c rs = true → is_marked ? c = false) → + ∀a,l. (a::l) = reverse ? (〈x,true〉::rs) → + t2 = rightof (FinProd alpha FinBool) a (l@ls)) ∧ + ∀l1,x0,l2. + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = l1@〈x0,true〉::l2 → + (x = x0 → + l2 = [ ] → (* test true but l2 is empty *) + t2 = rightof ? 〈x0,false〉 ((reverse ? l1)@〈x,true〉::ls)) ∧ + (x = x0 → + ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *) + 〈a,false〉::l1' = l1@[〈x0,false〉] → + l2 = 〈a0,b〉::l2' → + t2 = midtape ? (〈x,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧ + (x ≠ x0 →(* test false *) + t2 = midtape (FinProd … alpha FinBool) ((reverse ? l1)@〈x,true〉::ls) 〈x0,false〉 l2)) ∧ (〈x,true〉 ≠ c → RelseM t1 t2). lemma dec_marked: ∀alpha,rs. @@ -740,7 +814,9 @@ lemma dec_marked: ∀alpha,rs. |#Hnall %2 @(not_to_not … Hnall) #Hall #c #memc @Hall @memb_cons // ] qed. - + +(* axiom daemon:∀P:Prop.P. *) + lemma sem_comp_step_subcase : ∀alpha,c,elseM,RelseM. Realize ? elseM RelseM → @@ -751,43 +827,101 @@ cases (sem_if ? (test_char ? (λx.x == c)) … tc_true (sem_test_char ? (λx.x == c)) (sem_seq ????? (sem_move_r …) (sem_seq ????? (sem_adv_to_mark_r ? (is_marked alpha)) - (sem_match_and_adv ? (λx.x == c)))) Helse intape) + (sem_match_and_adv_full ? (λx.x == c)))) Helse intape) #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop cases HR -HR - [* #ta * whd in ⊢ (%→?); #Hta * #tb * whd in ⊢ (%→?); #Htb - * #tc * whd in ⊢ (%→?); #Htc whd in ⊢ (%→?); #Houtc - #l0 #x #rs #Hintape cases (true_or_false (〈x,true〉==c)) #Hc - [%1 #_ cases (dec_marked ? rs) #Hdec + [* #ta * whd in ⊢ (%→?); * * #cin * #Hcin #Hcintrue #Hta * #tb * whd in ⊢ (%→?); #Htb + * #tc * whd in ⊢ (%→?); #Htc * whd in ⊢ (%→%→?); #Houtc #Houtc1 + #ls #x #rs #Hintape >Hintape in Hcin; whd in ⊢ ((??%?)→?); #H destruct (H) % + [#_ cases (dec_marked ? rs) #Hdec [% [#_ #a #l1 - >Hintape in Hta; * #_(* #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) - #Hta lapply (proj2 … Htb … Hta) -Htb -Hta cases rs - [whd in ⊢ ((???%)→?); #Htb >Htb in Htc; #Htc - lapply (proj1 ?? Htc (refl …)) -Htc #Htc Hintape in Hta; #Hta + lapply (proj2 ?? Htb … Hta) -Htb -Hta cases rs in Hdec; + (* by cases on rs *) + [#_ whd in ⊢ ((???%)→?); #Htb >Htb in Htc; #Htc + lapply (proj1 ?? Htc (refl …)) -Htc #Htc Htb in Htc; #Htc + >reverse_cons >reverse_cons #Hl1 + cases (proj2 ?? Htc … (refl …)) + [* >(Hdec …) [ #Hfalse destruct(Hfalse) ] @memb_hd + |* #_ -Htc #Htc cut (∃l2.l1 = l2@[〈x,true〉]) + [generalize in match Hl1; -Hl1 <(reverse_reverse … l1) + cases (reverse ? l1) + [#Hl1 cut ([a]=〈x,true〉::r0::rs0) + [ <(reverse_reverse … (〈x,true〉::r0::rs0)) + >reverse_cons >reverse_cons reverse_cons #Heq + lapply (append_l2_injective_r ? (a::reverse ? l10) ???? Heq) // + #Ha0 destruct(Ha0) /2/ ] + |* #l2 #Hl2 >Hl2 in Hl1; #Hl1 + lapply (append_l1_injective_r ? (a::l2) … Hl1) // -Hl1 #Hl1 + >reverse_cons in Htc; #Htc lapply (Htc … (sym_eq … Hl1)) + [ #x0 #Hmemb @Hdec @memb_cons @Hmemb ] + -Htc #Htc >Htc in Houtc1; #Houtc1 >associative_append @Houtc1 % + ] + ] + ] + |#l1 #x0 #l2 #_ #Hrs @False_ind @(absurd ?? not_eq_true_false) change with (is_marked ? 〈x0,true〉) in match true; - @Hdec >Hrs @memb_cons @memb_append_l2 @memb_hd + @Hdec >Hrs @memb_append_l2 @memb_hd ] |% [#H @False_ind @(absurd …H Hdec)] - #a #l1 #x0 #a0 #l2 #Hl1 #Hrs >Hrs in Hintape; #Hintape - >Hintape in Hta; * #_(* #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) - #Hta lapply (proj2 … Htb … Hta) -Htb -Hta - whd in match (mk_tape ????); #Htb cases Htc -Htc #_ #Htc - cases (Htc … Htb) [ * #Hfalse normalize in Hfalse; destruct (Hfalse) ] - -Htc * #_ #Htc lapply (Htc l1 〈x0,true〉 (〈a0,false〉::l2) (refl ??) (refl ??) Hl1) - -Htc #Htc cases (Houtc ???????? Htc) -Houtc - [* #Hx0 #Houtc %1 #Hx >Houtc >reverse_reverse % - |* #Hx0 #Houtc %2 #_ >Houtc % - |#x #membx @Hl1 <(reverse_reverse …l1) @memb_reverse @membx + (* by cases on l1 *) * + [#x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape + >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) + #Hta lapply (proj2 … Htb … Hta) -Htb -Hta + whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc + #_ #Htc cases (Htc … Htb) -Htc + [2: * * #Hfalse normalize in Hfalse; destruct (Hfalse) ] + * * #Htc >Htb in Htc; -Htb #Htc cases (Houtc … Htc) -Houtc * + #H1 #H2 #H3 cases (true_or_false (x==x0)) #eqxx0 + [>(\P eqxx0) % [2: #H @False_ind /2/] % + [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0) [% | @Hcintrue] + |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes destruct (Hdes) + #Hl2 @(H3 … Hdec … Hl2) <(\P eqxx0) [@Hcintrue | % | @reverse_single] + ] + |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)] + #_ @H1 @(\bf ?) @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue) + #Hdes destruct (Hdes) % + ] + |#l1hd #l1tl #x0 #l2 #Hdec normalize in ⊢ (%→?); #Hrs >Hrs in Hintape; #Hintape + >Hintape in Hta; (* * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) #Hx *) + #Hta lapply (proj2 … Htb … Hta) -Htb -Hta + whd in match (mk_tape ????); whd in match (tail ??); #Htb cases Htc -Htc + #_ #Htc cases (Htc … Htb) -Htc + [* #Hfalse @False_ind >(Hdec … (memb_hd …)) in Hfalse; #H destruct] + * * #_ #Htc lapply (Htc … (refl …) (refl …) ?) -Htc + [#x1 #membx1 @Hdec @memb_cons @membx1] #Htc + cases (Houtc … Htc) -Houtc * + #H1 #H2 #H3 #_ cases (true_or_false (x==x0)) #eqxx0 + [>(\P eqxx0) % [2: #H @False_ind /2/] % + [#_ #Hl2 >(H2 … Hl2) <(\P eqxx0) + [>reverse_cons >associative_append % | @Hcintrue] + |#_ #a #a0 #b #l1' #l2' normalize in ⊢ (%→?); #Hdes (* destruct (Hdes) *) + #Hl2 @(H3 ?????? (reverse … (l1hd::l1tl)) … Hl2) <(\P eqxx0) + [@Hcintrue + |>reverse_cons >associative_append % + |#c0 #memc @Hdec <(reverse_reverse ? (l1hd::l1tl)) @memb_reverse @memc + |>Hdes >reverse_cons >reverse_reverse >(\P eqxx0) % + ] + ] + |% [% #eqx @False_ind lapply (\Pf eqxx0) #Habs @(absurd … eqx Habs)] + #_ >reverse_cons >associative_append @H1 @(\bf ?) + @(not_to_not ??? (\Pf eqxx0)) <(\P Hcintrue) #Hdes + destruct (Hdes) % + ] ] ] - |%2 >Hintape in Hta; * * #x1 * whd in ⊢ ((??%?)→?); #H destruct (H) - >Hc #H destruct (H) + |>(\P Hcintrue) * #Hfalse @False_ind @Hfalse % ] - |* #ta * whd in ⊢ (%→?); * #Hc #Hta #Helse #ls #c0 #rs #Hintape %2 - #_ >Hintape in Hta; #Hta Hintape in Hcur; #Hcur lapply (Hcur ? (refl …)) -Hcur #Hc % + [ #Hfalse >Hfalse in Hc; #Hc cases (\Pf Hc) #Hc @False_ind @Hc % + | -Hc #Hc Hfa normalize in ⊢ (%→?); #H destruct - [% [@(ex_intro … []) % normalize [% % | #x @False_ind] + [% [@(ex_intro … []) % normalize [% % | #x #H destruct] |#a1 #tl1 #H destruct (H) //] |cases (Hind (a::acc) res1 res2 H) * #l1 * * #Hres1 #Htl #Hfalse #Htrue % [2:@Htrue] @(ex_intro … (l1@[a])) % [% [>associative_append @Hres1 | >reverse_append H //| @False_ind] + |#x #Hmemx cases (memb_append ???? Hmemx) + [@Hfalse | #H >(memb_single … H) //] ] ] ] @@ -856,40 +990,64 @@ qed. axiom mem_reverse: ∀A,l,x. mem A x (reverse ? l) → mem A x l. -lemma split_on_spec_ex: ∀A,l,f.∃l1,l2. - l1@l2 = l ∧ (∀x:A. mem ? x l1 → f x = false) ∧ +lemma split_on_spec_ex: ∀A:DeqSet.∀l,f.∃l1,l2. + l1@l2 = l ∧ (∀x:A. memb ? x l1 = true → f x = false) ∧ ∀a,tl. l2 = a::tl → f a = true. #A #l #f @(ex_intro … (reverse … (\fst (split_on A l f [])))) @(ex_intro … (\snd (split_on A l f []))) cases (split_on_spec A l f [ ] ?? (eq_pair_fst_snd …)) * #l1 * * >append_nil #Hl1 >Hl1 #Hl #Hfalse #Htrue % - [% [@Hl|#x #memx @Hfalse @mem_reverse //] | @Htrue] + [% [@Hl|#x #memx @Hfalse <(reverse_reverse … l1) @memb_reverse //] | @Htrue] qed. +(* versione esistenziale *) + definition R_comp_step_true ≝ λt1,t2. - ∃l0,c,a,l1,c0,l1',a0,l2. - t1 = midtape (FinProd … FSUnialpha FinBool) - l0 〈c,true〉 (l1@〈c0,true〉::〈a0,false〉::l2) ∧ - l1@[〈c0,false〉] = 〈a,false〉::l1' ∧ - (∀c.memb ? c l1 = true → is_marked ? c = false) ∧ - (bit_or_null c = true → c0 = c → - t2 = midtape ? (〈c,false〉::l0) 〈a,true〉 (l1'@〈c0,false〉::〈a0,true〉::l2)) ∧ - (bit_or_null c = true → c0 ≠ c → - t2 = midtape (FinProd … FSUnialpha FinBool) - (reverse ? l1@〈a,false〉::〈c,true〉::l0) 〈c0,false〉 (〈a0,false〉::l2)) ∧ - (bit_or_null c = false → - t2 = midtape ? l0 〈c,false〉 (〈a,false〉::l1@〈c0,true〉::〈a0,false〉::l2)). + ∃ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls 〈c,true〉 rs ∧ + ((* bit_or_null c = false *) + (bit_or_null c = false → t2 = midtape ? ls 〈c,false〉 rs) ∧ + (* no marks in rs *) + (bit_or_null c = true → + (∀c.memb ? c rs = true → is_marked ? c = false) → + ∀a,l. (a::l) = reverse ? (〈c,true〉::rs) → + t2 = rightof (FinProd FSUnialpha FinBool) a (l@ls)) ∧ + (∀l1,c0,l2. + bit_or_null c = true → + (∀c.memb ? c l1 = true → is_marked ? c = false) → + rs = l1@〈c0,true〉::l2 → + (c = c0 → + l2 = [ ] → (* test true but l2 is empty *) + t2 = rightof ? 〈c0,false〉 ((reverse ? l1)@〈c,true〉::ls)) ∧ + (c = c0 → + ∀a,a0,b,l1',l2'. (* test true and l2 is not empty *) + 〈a,false〉::l1' = l1@[〈c0,false〉] → + l2 = 〈a0,b〉::l2' → + t2 = midtape ? (〈c,false〉::ls) 〈a,true〉 (l1'@〈a0,true〉::l2')) ∧ + (c ≠ c0 →(* test false *) + t2 = midtape (FinProd … FSUnialpha FinBool) + ((reverse ? l1)@〈c,true〉::ls) 〈c0,false〉 l2))). definition R_comp_step_false ≝ λt1,t2. ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → is_marked ? c = false ∧ t2 = t1. -(* lemma is_marked_to_exists: ∀alpha,c. is_marked alpha c = true → ∃c'. c = 〈c',true〉. -#alpha * c *) +#alpha * #c * [#_ @(ex_intro … c) //| normalize #H destruct] +qed. +lemma exists_current: ∀alpha,c,t. + current alpha t = Some alpha c → ∃ls,rs. t= midtape ? ls c rs. +#alpha #c * + [whd in ⊢ (??%?→?); #H destruct + |#a #l whd in ⊢ (??%?→?); #H destruct + |#a #l whd in ⊢ (??%?→?); #H destruct + |#ls #c1 #rs whd in ⊢ (??%?→?); #H destruct + @(ex_intro … ls) @(ex_intro … rs) // + ] +qed. + lemma sem_comp_step : accRealize ? comp_step (inr … (inl … (inr … start_nop))) R_comp_step_true R_comp_step_false. @@ -899,149 +1057,60 @@ lemma sem_comp_step : (sem_comp_step_subcase FSUnialpha 〈null,true〉 ?? (sem_clear_mark …)))) (sem_nop …) …) -[#intape #outtape #midtape * * * #c #b * #Hcurrent -whd in ⊢ ((??%?)→?); #Hb #Hmidtape >Hmidtape -Hmidtape - cases (current_to_midtape … Hcurrent) #ls * #rs >Hb #Hintape >Hintape -Hb - whd in ⊢ (%→?); #Htapea lapply (Htapea … (refl …)) -Htapea - cases (split_on_spec_ex ? rs (is_marked ?)) #l1 * #l2 * * #Hrs #Hl1 #Hl2 - cases (true_or_false (c == bit false)) - [(* c = bit false *) #Hc * [2: * >(\P Hc) * #H @False_ind @H //] - * #_ #Hout whd - cases (split_on_spec - - -[ #Hstate lapply (H1 Hstate) -H1 -Hstate -H2 * - #ta * whd in ⊢ (%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; * * - cases c in Hintape; #c' #b #Hintape #x * whd in ⊢ (??%?→?); #H destruct (H) - whd in ⊢ (??%?→?); #Hb >Hb #Hta @(ex_intro ?? c') % // - cases (Hright … Hta) - [ * #Hc' #H1 % % [destruct (Hc') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] - ] - | * #Hc #Helse1 cases (Helse1 … Hta) - [ * #Hc' #H1 % % [destruct (Hc') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] +[#intape #outape #ta #Hta #Htb cases Hta * #cm * #Hcur + cases (exists_current … Hcur) #ls * #rs #Hintape #cmark + cases (is_marked_to_exists … cmark) #c #Hcm + >Hintape >Hcm -Hintape -Hcm #Hta + @(ex_intro … ls) @(ex_intro … c) @(ex_intro …rs) % [//] lapply Hta -Hta + (* #ls #c #rs #Hintape whd in Hta; + >Hintape in Hta; * #_ -Hintape forse non serve *) + cases (true_or_false (c==bit false)) #Hc + [>(\P Hc) #Hta % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 ?? (proj1 ?? (Htb … Hta) (refl …))) ] - | * #Hc' #Helse2 cases (Helse2 … Hta) - [ * #Hc'' #H1 % % [destruct (Hc'') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] - ] - | * #Hc'' whd in ⊢ (%→?); #Helse3 %2 % - [ generalize in match Hc''; generalize in match Hc'; generalize in match Hc; - cases c' - [ * [ #_ #Hfalse @False_ind @(absurd ?? Hfalse) % - | #Hfalse @False_ind @(absurd ?? Hfalse) % ] - | #_ #_ #Hfalse @False_ind @(absurd ?? Hfalse) % - |*: #_ #_ #_ % ] - | @(Helse3 … Hta) + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (Htb … Hta) (refl …))) + ] + |cases (true_or_false (c==bit true)) #Hc1 + [>(\P Hc1) #Hta + cut (〈bit true, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) (refl …))) ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (Htb … Hta) Hneq … Hta)(refl …))) ] - ] - ] -| #Hstate lapply (H2 Hstate) -H1 -Hstate -H2 * - #ta * whd in ⊢ (%→%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; * #Hc #Hta % [@Hc % | >Hright //] -] -qed. -definition R_comp_step_true ≝ - λt1,t2. - ∀l0,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) l0 c rs → - ∃c'. c = 〈c',true〉 ∧ - ((bit_or_null c' = true ∧ - ∀a,l1,c0,a0,l2. - rs = 〈a,false〉::l1@〈c0,true〉::〈a0,false〉::l2 → - (∀c.memb ? c l1 = true → is_marked ? c = false) → - (c0 = c' ∧ - t2 = midtape ? (〈c',false〉::l0) 〈a,true〉 (l1@〈c0,false〉::〈a0,true〉::l2)) ∨ - (c0 ≠ c' ∧ - t2 = midtape (FinProd … FSUnialpha FinBool) - (reverse ? l1@〈a,false〉::〈c',true〉::l0) 〈c0,false〉 (〈a0,false〉::l2))) ∨ - (bit_or_null c' = false ∧ t2 = midtape ? l0 〈c',false〉 rs)). - -definition R_comp_step_false ≝ - λt1,t2. - ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → - is_marked ? c = false ∧ t2 = t1. - -lemma sem_comp_step : - 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 ?)) - (sem_comp_step_subcase FSUnialpha 〈bit false,true〉 ?? - (sem_comp_step_subcase FSUnialpha 〈bit true,true〉 ?? - (sem_comp_step_subcase FSUnialpha 〈null,true〉 ?? - (sem_clear_mark …)))) - (sem_nop …) intape) -#k * #outc * * #Hloop #H1 #H2 -@(ex_intro ?? k) @(ex_intro ?? outc) % -[ % [@Hloop ] ] -Hloop -[ #Hstate lapply (H1 Hstate) -H1 -Hstate -H2 * - #ta * whd in ⊢ (%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; * * - cases c in Hintape; #c' #b #Hintape #x * whd in ⊢ (??%?→?); #H destruct (H) - whd in ⊢ (??%?→?); #Hb >Hb #Hta @(ex_intro ?? c') % // - cases (Hright … Hta) - [ * #Hc' #H1 % % [destruct (Hc') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] - ] - | * #Hc #Helse1 cases (Helse1 … Hta) - [ * #Hc' #H1 % % [destruct (Hc') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] - ] - | * #Hc' #Helse2 cases (Helse2 … Hta) - [ * #Hc'' #H1 % % [destruct (Hc'') % ] - #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1 - cases (H1 … Hl1 Hrs) - [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc - | * #Hneq #Houtc %2 % - [ @sym_not_eq // - | @Houtc ] + |cases (true_or_false (c==null)) #Hc2 + [>(\P Hc2) #Hta + cut (〈null, true〉 ≠ 〈bit false, true〉) [% #Hdes destruct] #Hneq + cut (〈null, true〉 ≠ 〈bit true, true〉) [% #Hdes destruct] #Hneq1 % + [%[whd in ⊢ ((??%?)→?); #Hdes destruct + |#Hc @(proj1 … (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) + ] + |#l1 #c0 #l2 #Hc @(proj2 ?? (proj1 ?? (proj2 ?? (proj2 ?? (Htb … Hta) Hneq … Hta) Hneq1 … Hta) (refl …))) ] - | * #Hc'' whd in ⊢ (%→?); #Helse3 %2 % - [ generalize in match Hc''; generalize in match Hc'; generalize in match Hc; - cases c' - [ * [ #_ #Hfalse @False_ind @(absurd ?? Hfalse) % - | #Hfalse @False_ind @(absurd ?? Hfalse) % ] - | #_ #_ #Hfalse @False_ind @(absurd ?? Hfalse) % - |*: #_ #_ #_ % ] - | @(Helse3 … Hta) + |#Hta cut (bit_or_null c = false) + [lapply Hc; lapply Hc1; lapply Hc2 -Hc -Hc1 -Hc2 + cases c normalize [* normalize /2/] /2/] #Hcut % + [%[cases (Htb … Hta) #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc; destruct] #_ -Htb #Htb + cases (Htb … Hta) [2: % #H destruct (H) normalize in Hc1; destruct] #_ -Htb #Htb + lapply (Htb ?) [% #H destruct (H) normalize in Hc2; destruct] + * #_ #Houttape #_ @(Houttape … Hta) + |>Hcut #H destruct + ] + |#l1 #c0 #l2 >Hcut #H destruct ] ] ] ] -| #Hstate lapply (H2 Hstate) -H1 -Hstate -H2 * - #ta * whd in ⊢ (%→%→?); #Hleft #Hright #ls #c #rs #Hintape - >Hintape in Hleft; * #Hc #Hta % [@Hc % | >Hright //] +|#intape #outape #ta #Hta #Htb #ls #c #rs #Hintape + >Hintape in Hta; whd in ⊢ (%→?); * #Hmark #Hta % [@Hmark //] + whd in Htb; >Htb // ] qed. +(* compare *) + definition compare ≝ whileTM ? comp_step (inr … (inl … (inr … start_nop))). @@ -1091,11 +1160,19 @@ RIFIUTO: c ≠ d t2 = midtape ? l0 〈grid,false〉 (l1@〈comma,true〉::l2)) ∨ (b = bit x ∧ b = c ∧ bs = b0s *) + +definition list_cases2: ∀A.∀P:list A →list A →Prop.∀l1,l2. |l1| = |l2| → +P [ ] [ ] → (∀a1,a2:A.∀tl1,tl2. |tl1| = |tl2| → P (a1::tl1) (a2::tl2)) → P l1 l2. +#A #P #l1 #l2 #Hlen lapply Hlen @(list_ind2 … Hlen) // +#tl1 #tl2 #hd1 #hd2 #Hind normalize #HlenS #H1 #H2 @H2 // +qed. + definition R_compare := λt1,t2. ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → (∀c'.bit_or_null c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧ (∀c'. c = 〈c',false〉 → t2 = t1) ∧ +(* forse manca il caso no marks in rs *) ∀b,b0,bs,b0s,l1,l2. |bs| = |b0s| → (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → bit_or_null (\fst c) = true) → @@ -1133,103 +1210,179 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] cases (Rfalse … Htapea) -Rfalse >Hc whd in ⊢ (??%?→?);#Hfalse destruct (Hfalse) ] | #tapea #tapeb #tapec #Hleft #Hright #IH #Htapec lapply (IH Htapec) -Htapec -IH #IH - whd in Hleft; #ls #c #rs #Htapea cases (Hleft … Htapea) -Hleft - #c' * #Hc >Hc cases (true_or_false (bit_or_null c')) #Hc' - [2: * - [ * >Hc' #H @False_ind destruct (H) - | * #_ #Htapeb cases (IH … Htapeb) * #_ #H #_ % - [% - [#c1 #Hc1 #Heqc destruct (Heqc) Hc' #Hfalse @False_ind destruct (Hfalse) + whd in Hleft; #ls #c #rs #Htapea cases Hleft -Hleft + #ls0 * #c' * #rs0 * >Htapea #Hdes destruct (Hdes) * * + cases (true_or_false (bit_or_null c')) #Hc' + [2: #Htapeb lapply (Htapeb Hc') -Htapeb #Htapeb #_ #_ % + [%[#c1 #Hc1 #Heqc destruct (Heqc) + cases (IH … Htapeb) * #_ #H #_ Hc' #Hfalse @False_ind destruct (Hfalse) ] - |#Hleft % + |#_ (* no marks in rs ??? *) #_ #Hleft % [ % [ #c'' #Hc'' #Heq destruct (Heq) >Hc'' in Hc'; #H destruct (H) | #c0 #Hfalse destruct (Hfalse) ] |#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_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 #_ #_ - % % - [ >Heqb >Hbs >Hb0s % - | >Hbs >Hb0s @IH % - ] - |* #Hneqb #Htapeb %2 - @(ex_intro … [ ]) @(ex_intro … b) - @(ex_intro … b0) @(ex_intro … [ ]) - @(ex_intro … [ ]) % - [ % [ % [@sym_not_eq //| >Hbs %] | >Hb0s %] - | cases (IH … Htapeb) -IH * #_ #IH #_ >(IH ? (refl ??)) - @Htapeb + #Heq destruct (Heq) #_ >append_cons; (memb_single … memc1) %] + |@Hl1 @memc1 + ] + |* (* manca il caso in cui dopo una parte uguale il + secondo nastro finisca ... ???? *) + #_ cases (true_or_false (b==b0)) #eqbb0 + [2: #_ #Htapeb %2 lapply (Htapeb … (\Pf eqbb0)) -Htapeb #Htapeb + cases (IH … Htapeb) * #_ #Hout #_ + @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0) + @(ex_intro … bs) @(ex_intro … b0s) % + [%[%[@(\Pf eqbb0) | %] | %] + |>(Hout … (refl …)) -Hout >Htapeb @eq_f3 [2,3:%] + >reverse_append >reverse_append >associative_append + >associative_append % + ] + |lapply Hbs1 lapply Hb0s1 lapply Hbs2 lapply Hb0s2 lapply Hrs + -Hbs1 -Hb0s1 -Hbs2 -Hb0s2 -Hrs + @(list_cases2 … Hlen) + [#Hrs #_ #_ #_ #_ >associative_append >associative_append #Htapeb #_ + lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb + cases (IH … Htapeb) -IH * #Hout #_ #_ %1 % + [>(\P eqbb0) % + |>(Hout grid (refl …) (refl …)) @eq_f + normalize >associative_append % ] - | @Hl1 ] - | * #b' #bitb' * #b0' #bitb0' #bs' #b0s' #Hbs #Hb0s - generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?); - cut (bit_or_null b' = true ∧ bit_or_null b0' = true ∧ - bitb' = false ∧ bitb0' = false) - [ % [ % [ % [ >Hbs in Hbs1; #Hbs1 @(Hbs1 〈b',bitb'〉) @memb_hd - | >Hb0s in Hb0s1; #Hb0s1 @(Hb0s1 〈b0',bitb0'〉) @memb_hd ] - | >Hbs in Hbs2; #Hbs2 @(Hbs2 〈b',bitb'〉) @memb_hd ] - | >Hb0s in Hb0s2; #Hb0s2 @(Hb0s2 〈b0',bitb0'〉) @memb_hd ] - | * * * #Ha #Hb #Hc #Hd >Hc >Hd - #Hrs #Hleft - cases (Hleft b' (bs'@〈grid,false〉::l1) b0 b0' - (b0s'@〈comma,false〉::l2) ??) -Hleft - [ 3: >Hrs normalize @eq_f >associative_append % - | * #Hb0 #Htapeb cases (IH …Htapeb) -IH * #_ #_ #IH - cases (IH b' b0' bs' b0s' (l1@[〈b0,false〉]) l2 ??????? Ha ?) -IH - [ * #Heq #Houtc % % - [ >Hb0 @eq_f >Hbs in Heq; >Hb0s in ⊢ (%→?); #Heq - destruct (Heq) >Hb0s >Hc >Hd % - | >Houtc >Hbs >Hb0s >Hc >Hd >reverse_cons >associative_append - >associative_append % + |* #a1 #ba1 * #a2 #ba2 #tl1 #tl2 #HlenS #Hrs #Hb0s2 #Hbs2 #Hb0s1 #Hbs1 + cut (ba1 = false) [@(Hbs2 〈a1,ba1〉) @memb_hd] #Hba1 >Hba1 + >associative_append >associative_append #Htapeb #_ + lapply (Htapeb … (\P eqbb0) … (refl …) (refl …)) -Htapeb #Htapeb + cases (IH … Htapeb) -IH * #_ #_ + cut (ba2=false) [@(Hb0s2 〈a2,ba2〉) @memb_hd] #Hba2 >Hba2 + #IH cases(IH a1 a2 ?? (l1@[〈b0,false〉]) l2 HlenS ????? (refl …) ??) + [3:#x #memx @Hbs1 @memb_cons @memx + |4:#x #memx @Hb0s1 @memb_cons @memx + |5:#x #memx @Hbs2 @memb_cons @memx + |6:#x #memx @Hb0s2 @memb_cons @memx + |7:#x #memx cases (memb_append …memx) -memx #memx + [@Hl1 @memx | >(memb_single … memx) %] + |8:@(Hbs1 〈a1,ba1〉) @memb_hd + |9: >associative_append >associative_append % + |-IH -Hbs1 -Hb0s1 -Hbs2 -Hrs * + #Ha1a2 #Houtc %1 % + [>(\P eqbb0) @eq_f destruct (Ha1a2) % + |>Houtc @eq_f3 + [>reverse_cons >associative_append % + |% + |>associative_append % + ] ] - | * #la * #c' * #d' * #lb * #lc * * * #H1 #H2 #H3 #H4 %2 - @(ex_intro … (〈b,false〉::la)) @(ex_intro … c') @(ex_intro … d') - @(ex_intro … lb) @(ex_intro … lc) - % [ % [ % // >Hbs >Hc >H2 % | >Hb0s >Hd >H3 >Hb0 % ] - | >H4 >Hbs >Hb0s >Hc >Hd >Hb0 >reverse_append - >reverse_cons >reverse_cons - >associative_append >associative_append - >associative_append >associative_append % + |-IH -Hbs1 -Hb0s1 -Hbs2 -Hrs * + #la * #c' * #d' * #lb * #lc * * * + #Hcd #H1 #H2 #Houtc %2 + @(ex_intro … (〈b,false〉::la)) @(ex_intro … c') @(ex_intro … d') + @(ex_intro … lb) @(ex_intro … lc) % + [%[%[@Hcd | >H1 %] |>(\P eqbb0) >Hba2 >H2 %] + |>Houtc @eq_f3 + [>(\P eqbb0) >reverse_append >reverse_cons + >reverse_cons >associative_append >associative_append + >associative_append >associative_append % + |% + |% ] - | generalize in match Hlen; >Hbs >Hb0s - normalize #Hlen destruct (Hlen) @e0 - | #c0 #Hc0 @Hbs1 >Hbs @memb_cons // - | #c0 #Hc0 @Hb0s1 >Hb0s @memb_cons // - | #c0 #Hc0 @Hbs2 >Hbs @memb_cons // - | #c0 #Hc0 @Hb0s2 >Hb0s @memb_cons // - | #c0 #Hc0 cases (memb_append … Hc0) - [ @Hl1 | #Hc0' >(memb_single … Hc0') % ] - | % - | >associative_append >associative_append % ] - | * #Hneq #Htapeb %2 - @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0) - @(ex_intro … bs) @(ex_intro … b0s) % - [ % // % // @sym_not_eq // - | >Hbs >Hb0s >Hc >Hd >reverse_cons >associative_append - >reverse_append in Htapeb; >reverse_cons - >associative_append >associative_append - #Htapeb Hbs @memb_cons @Hyp - | cases (orb_true_l … Hyp) - [ #Hyp2 >(\P Hyp2) % - | @Hl1 - ] - ] - ] -]]]]] -qed. + ] + ] + ] + ] + ] + ] + ] +] +qed. + +lemma WF_cst_niltape: + WF ? (inv ? R_comp_step_true) (niltape (FinProd FSUnialpha FinBool)). +@wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct +qed. + +lemma WF_cst_rightof: + ∀a,ls. WF ? (inv ? R_comp_step_true) (rightof (FinProd FSUnialpha FinBool) a ls). +#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct +qed. + +lemma WF_cst_leftof: + ∀a,ls. WF ? (inv ? R_comp_step_true) (leftof (FinProd FSUnialpha FinBool) a ls). +#a #ls @wf #t1 whd in ⊢ (%→?); * #ls * #c * #rs * #H destruct +qed. + +lemma WF_cst_midtape_false: + ∀ls,c,rs. WF ? (inv ? R_comp_step_true) + (midtape (FinProd … FSUnialpha FinBool) ls 〈c,false〉 rs). +#ls #c #rs @wf #t1 whd in ⊢ (%→?); * #ls' * #c' * #rs' * #H destruct +qed. + +(* da spostare *) +lemma not_nil_to_exists:∀A.∀l: list A. l ≠ [ ] → + ∃a,tl. a::tl = l. + #A * [* #H @False_ind @H // | #a #tl #_ @(ex_intro … a) @(ex_intro … tl) //] + qed. + +lemma terminate_compare: + ∀t. Terminate ? compare t. +#t @(terminate_while … sem_comp_step) [%] +cases t // #ls * #c * // +#rs +(* we cannot proceed by structural induction on the right tape, + since compare moves the marks! *) +cut (∃len. |rs| = len) [/2/] +* #len lapply rs lapply c lapply ls -ls -c -rs elim len + [#ls #c #rs #Hlen >(lenght_to_nil … Hlen) @wf #t1 whd in ⊢ (%→?); * #ls0 * #c0 * #rs0 * #Hmid destruct (Hmid) + * * #H1 #H2 #_ cases (true_or_false (bit_or_null c0)) #Hc0 + [>(H2 Hc0 … (refl …)) // #x whd in ⊢ ((??%?)→?); #Hdes destruct + |>(H1 Hc0) // + ] + |-len #len #Hind #ls #c #rs #Hlen @wf #t1 whd in ⊢ (%→?); * #ls0 * #c0 * #rs0 * #Hmid destruct (Hmid) + * * #H1 #H2 #H3 cases (true_or_false (bit_or_null c0)) #Hc0 + [-H1 cases (split_on_spec_ex ? rs0 (is_marked ?)) #rs1 * #rs2 + cases rs2 + [(* no marks in right tape *) + * * >append_nil #H >H -H #Hmarks #_ + cases (not_nil_to_exists ? (reverse (FSUnialpha×bool) (〈c0,true〉::rs0)) ?) + [2: % >reverse_cons #H cases (nil_to_nil … H) #_ #H1 destruct] + #a0 * #tl #H4 >(H2 Hc0 Hmarks a0 tl H4) // + |(* the first marked element is a0 *) + * #a0 #a0b #rs3 * * #H4 #H5 #H6 lapply (H3 ? a0 rs3 … Hc0 H5 ?) + [

(Ht1 (\P eqc0a0) (refl …)) // + |(* a1 will be marked *) + cases (not_nil_to_exists ? (rs1@[〈a0,false〉]) ?) + [2: % #H cases (nil_to_nil … H) #_ #H1 destruct] + * #a2 #a2b * #tl2 #H7 * #a1 #a1b #rs4 #H4 #_ #Ht1 #_ + cut (a2b =false) + [lapply (memb_hd ? 〈a2,a2b〉 tl2) >H7 #mema2 + cases (memb_append … mema2) + [@H5 |#H lapply(memb_single … H) #H2 destruct %] + ] + #Ha2b >Ha2b in H7; #H7 + >(Ht1 (\P eqc0a0) … H7 (refl …)) @Hind -Hind -Ht1 -Ha2b -H2 -H3 -H5 -H6 +

length_append normalize length_append normalize <(injective_S … Hlen1) @eq_f2 // + cut (|〈a2,false〉::tl2|=|rs1@[〈a0,false〉]|) [>H7 %] + >length_append normalize (Ht1 (\Pf eqc0a0)) // + ] + ] + ] + |>(H1 Hc0) // + ] +qed. -axiom sem_compare : Realize ? compare R_compare. \ No newline at end of file +lemma sem_compare : Realize ? compare R_compare. +/2/ qed.