X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fre%2Fmoves.ma;h=c260a6d403dd34c60cc0697a2cf64568a1654b36;hb=927bda3b4b7fe5f521ae73eb008a746e8606a0b4;hp=b5449e4df179036e710c52c7458f786b1d7dde1f;hpb=d5d5925101dd773efb2f90136adc5d714a530cb9;p=helm.git diff --git a/matita/matita/lib/re/moves.ma b/matita/matita/lib/re/moves.ma index b5449e4df..c260a6d40 100644 --- a/matita/matita/lib/re/moves.ma +++ b/matita/matita/lib/re/moves.ma @@ -13,10 +13,11 @@ (**************************************************************************) include "re/re.ma". +include "basics/lists/listb.ma". -let rec move (S: Alpha) (x:S) (E: pitem S) on E : pre S ≝ +let rec move (S: DeqSet) (x:S) (E: pitem S) on E : pre S ≝ match E with - [ pz ⇒ 〈 ∅, false 〉 + [ pz ⇒ 〈 `∅, false 〉 | pe ⇒ 〈 ϵ, false 〉 | ps y ⇒ 〈 `y, false 〉 | pp y ⇒ 〈 `y, x == y 〉 @@ -24,27 +25,21 @@ let rec move (S: Alpha) (x:S) (E: pitem S) on E : pre S ≝ | pc e1 e2 ⇒ (move ? x e1) ⊙ (move ? x e2) | pk e ⇒ (move ? x e)^⊛ ]. -lemma move_plus: ∀S:Alpha.∀x:S.∀i1,i2:pitem S. +lemma move_plus: ∀S:DeqSet.∀x:S.∀i1,i2:pitem S. move S x (i1 + i2) = (move ? x i1) ⊕ (move ? x i2). // qed. -lemma move_cat: ∀S:Alpha.∀x:S.∀i1,i2:pitem S. +lemma move_cat: ∀S:DeqSet.∀x:S.∀i1,i2:pitem S. move S x (i1 · i2) = (move ? x i1) ⊙ (move ? x i2). // qed. -lemma move_star: ∀S:Alpha.∀x:S.∀i:pitem S. +lemma move_star: ∀S:DeqSet.∀x:S.∀i:pitem S. move S x i^* = (move ? x i)^⊛. // qed. -lemma fst_eq : ∀A,B.∀a:A.∀b:B. \fst 〈a,b〉 = a. -// qed. - -lemma snd_eq : ∀A,B.∀a:A.∀b:B. \snd 〈a,b〉 = b. -// qed. - -definition pmove ≝ λS:Alpha.λx:S.λe:pre S. move ? x (\fst e). +definition pmove ≝ λS:DeqSet.λx:S.λe:pre S. move ? x (\fst e). -lemma pmove_def : ∀S:Alpha.∀x:S.∀i:pitem S.∀b. +lemma pmove_def : ∀S:DeqSet.∀x:S.∀i:pitem S.∀b. pmove ? x 〈i,b〉 = move ? x i. // qed. @@ -53,93 +48,66 @@ lemma eq_to_eq_hd: ∀A.∀l1,l2:list A.∀a,b. #A #l1 #l2 #a #b #H destruct // qed. -axiom same_kernel: ∀S:Alpha.∀a:S.∀i:pitem S. +lemma same_kernel: ∀S:DeqSet.∀a:S.∀i:pitem S. |\fst (move ? a i)| = |i|. -(* #S #a #i elim i // - [#i1 #i2 >move_cat - cases (move S a i1) #i11 #b1 >fst_eq #IH1 - cases (move S a i2) #i21 #b2 >fst_eq #IH2 - normalize *) - -axiom iff_trans:∀A,B,C. A ↔ B → B ↔ C → A ↔ C. -axiom iff_or_l: ∀A,B,C. A ↔ B → C ∨ A ↔ C ∨ B. -axiom iff_or_r: ∀A,B,C. A ↔ B → A ∨ C ↔ B ∨ C. - -axiom epsilon_in_star: ∀S.∀A:word S → Prop. A^* [ ]. +#S #a #i elim i // + [#i1 #i2 >move_cat #H1 #H2 whd in ⊢ (???%);

move_plus #H1 #H2 whd in ⊢ (???%);

H normalize - [>(proj1 … (eqb_true …) H) % - [* // #bot @False_ind //| #H1 destruct /2/] - |% [#bot @False_ind // - | #H1 destruct @(absurd ((a==a)=true)) - [>(proj2 … (eqb_true …) (refl …)) // | /2/] - ] + [>(\P H) % [* // #bot @False_ind //| #H1 destruct /2/] + |% [@False_ind |#H1 cases (\Pf H) #H2 @H2 destruct //] ] |#i1 #i2 #HI1 #HI2 #w >(sem_cat S i1 i2) >move_cat @iff_trans[|@sem_odot] >same_kernel >sem_cat_w - @iff_trans[||@(iff_or_l … (HI2 w))] @iff_or_r % - [* #w1 * #w2 * * #eqw #w1in #w2in @(ex_intro … (a::w1)) - @(ex_intro … w2) % // % normalize // cases (HI1 w1) /2/ - |* #w1 * #w2 * cases w1 - [* #_ #H @False_ind /2/ - |#x #w3 * #eqaw normalize in eqaw; destruct #w3in #w2in - @(ex_intro … w3) @(ex_intro … w2) % // % // cases (HI1 w3) /2/ - ] - ] + @iff_trans[||@(iff_or_l … (HI2 w))] @iff_or_r + @iff_trans[||@iff_sym @deriv_middot //] + @cat_ext_l @HI1 |#i1 #i2 #HI1 #HI2 #w >(sem_plus S i1 i2) >move_plus >sem_plus_w @iff_trans[|@sem_oplus] @iff_trans[|@iff_or_l [|@HI2]| @iff_or_r //] |#i1 #HI1 #w >move_star - @iff_trans[|@sem_ostar] >same_kernel >sem_star_w % - [* #w1 * #w2 * * #eqw #w1in #w2in - @(ex_intro … (a::w1)) @(ex_intro … w2) % // % normalize // - cases (HI1 w1 ) /2/ - |* #w1 * #w2 * cases w1 - [* #_ #H @False_ind /2/ - |#x #w3 * #eqaw normalize in eqaw; destruct #w3in #w2in - @(ex_intro … w3) @(ex_intro … w2) % // % // cases (HI1 w3) /2/ - ] - ] + @iff_trans[|@sem_ostar] >same_kernel >sem_star_w + @iff_trans[||@iff_sym @deriv_middot //] + @cat_ext_l @HI1 ] qed. notation > "x ↦* E" non associative with precedence 60 for @{moves ? $x $E}. -let rec moves (S : Alpha) w e on w : pre S ≝ +let rec moves (S : DeqSet) w e on w : pre S ≝ match w with [ nil ⇒ e | cons x w' ⇒ w' ↦* (move S x (\fst e))]. -lemma moves_empty: ∀S:Alpha.∀e:pre S. +lemma moves_empty: ∀S:DeqSet.∀e:pre S. moves ? [ ] e = e. // qed. -lemma moves_cons: ∀S:Alpha.∀a:S.∀w.∀e:pre S. +lemma moves_cons: ∀S:DeqSet.∀a:S.∀w.∀e:pre S. moves ? (a::w) e = moves ? w (move S a (\fst e)). // qed. -lemma not_epsilon_sem: ∀S:Alpha.∀a:S.∀w: word S. ∀e:pre S. +lemma not_epsilon_sem: ∀S:DeqSet.∀a:S.∀w: word S. ∀e:pre S. iff ((a::w) ∈ e) ((a::w) ∈ \fst e). -#S #a #w * #i #b >fst_eq cases b normalize +#S #a #w * #i #b cases b normalize [% /2/ * // #H destruct |% normalize /2/] qed. -lemma same_kernel_moves: ∀S:Alpha.∀w.∀e:pre S. +lemma same_kernel_moves: ∀S:DeqSet.∀w.∀e:pre S. |\fst (moves ? w e)| = |\fst e|. #S #w elim w // qed. -axiom iff_not: ∀A,B. (iff A B) →(iff (¬A) (¬B)). - -axiom iff_sym: ∀A,B. (iff A B) →(iff B A). - -theorem decidable_sem: ∀S:Alpha.∀w: word S. ∀e:pre S. +theorem decidable_sem: ∀S:DeqSet.∀w: word S. ∀e:pre S. (\snd (moves ? w e) = true) ↔ \sem{e} w. #S #w elim w [* #i #b >moves_empty cases b % /2/ @@ -153,7 +121,7 @@ lemma not_true_to_false: ∀b.b≠true → b =false. #b * cases b // #H @False_ind /2/ qed. -theorem equiv_sem: ∀S:Alpha.∀e1,e2:pre S. +theorem equiv_sem: ∀S:DeqSet.∀e1,e2:pre S. iff (\sem{e1} =1 \sem{e2}) (∀w.\snd (moves ? w e1) = \snd (moves ? w e2)). #S #e1 #e2 % [#same_sem #w @@ -164,29 +132,21 @@ theorem equiv_sem: ∀S:Alpha.∀e1,e2:pre S. |#H #w1 @iff_trans [||@decidable_sem] moves_cons >moves_cons // +qed. -definition in_moves ≝ λS:Alpha.λw.λe:pre S. \snd(w ↦* e). +definition in_moves ≝ λS:DeqSet.λw.λe:pre S. \snd(w ↦* e). -coinductive equiv (S:Alpha) : pre S → pre S → Prop ≝ +(* +coinductive equiv (S:DeqSet) : pre S → pre S → Prop ≝ mk_equiv: ∀e1,e2: pre S. \snd e1 = \snd e2 → (∀x. equiv S (move ? x (\fst e1)) (move ? x (\fst e2))) → equiv S e1 e2. - -definition beqb ≝ λb1,b2. - match b1 with - [ true ⇒ b2 - | false ⇒ notb b2 - ]. - -lemma beqb_ok: ∀b1,b2. iff (beqb b1 b2 = true) (b1 = b2). -#b1 #b2 cases b1 cases b2 normalize /2/ -qed. - -definition Bin ≝ mk_Alpha bool beqb beqb_ok. +*) let rec beqitem S (i1,i2: pitem S) on i1 ≝ match i1 with @@ -203,96 +163,63 @@ let rec beqitem S (i1,i2: pitem S) on i1 ≝ | pk i11 ⇒ match i2 with [ pk i21 ⇒ beqitem S i11 i21 | _ ⇒ false] ]. -axiom beqitem_ok: ∀S,i1,i2. iff (beqitem S i1 i2 = true) (i1 = i2). - -definition BinItem ≝ - mk_Alpha (pitem Bin) (beqitem Bin) (beqitem_ok Bin). - -definition beqpre ≝ λS:Alpha.λe1,e2:pre S. - beqitem S (\fst e1) (\fst e2) ∧ beqb (\snd e1) (\snd e2). - -definition beqpairs ≝ λS:Alpha.λp1,p2:(pre S)×(pre S). - beqpre S (\fst p1) (\fst p2) ∧ beqpre S (\snd p1) (\snd p2). - -axiom beqpairs_ok: ∀S,p1,p2. iff (beqpairs S p1 p2 = true) (p1 = p2). - -definition space ≝ λS.mk_Alpha ((pre S)×(pre S)) (beqpairs S) (beqpairs_ok S). - -let rec memb (S:Alpha) (x:S) (l: list S) on l ≝ - match l with - [ nil ⇒ false - | cons a tl ⇒ (a == x) || memb S x tl - ]. - -lemma memb_hd: ∀S,a,l. memb S a (a::l) = true. -#S #a #l normalize >(proj2 … (eqb_true S …) (refl S a)) // -qed. - -lemma memb_cons: ∀S,a,b,l. - memb S a l = true → memb S a (b::l) = true. -#S #a #b #l normalize cases (b==a) normalize // -qed. - -lemma memb_append: ∀S,a,l1,l2. -memb S a (l1@l2) = true → - memb S a l1= true ∨ memb S a l2 = true. -#S #a #l1 elim l1 normalize [/2/] #b #tl #Hind -#l2 cases (b==a) normalize /2/ -qed. - -lemma memb_append_l1: ∀S,a,l1,l2. - memb S a l1= true → memb S a (l1@l2) = true. -#S #a #l1 elim l1 normalize - [normalize #le #abs @False_ind /2/ - |#b #tl #Hind #l2 cases (b==a) normalize /2/ +lemma beqitem_true: ∀S,i1,i2. iff (beqitem S i1 i2 = true) (i1 = i2). +#S #i1 elim i1 + [#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct + |#i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % // normalize #H destruct + |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct + [>(\P H) // | @(\b (refl …))] + |#x #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % normalize #H destruct + [>(\P H) // | @(\b (refl …))] + |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % + normalize #H destruct + [cases (true_or_false (beqitem S i11 i21)) #H1 + [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H + |>H1 in H; normalize #abs @False_ind /2/ + ] + |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) // + ] + |#i11 #i12 #Hind1 #Hind2 #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i3] % + normalize #H destruct + [cases (true_or_false (beqitem S i11 i21)) #H1 + [>(proj1 … (Hind1 i21) H1) >(proj1 … (Hind2 i22)) // >H1 in H; #H @H + |>H1 in H; normalize #abs @False_ind /2/ + ] + |>(proj2 … (Hind1 i21) (refl …)) >(proj2 … (Hind2 i22) (refl …)) // + ] + |#i3 #Hind #i2 cases i2 [||#a|#a|#i21 #i22| #i21 #i22|#i4] % + normalize #H destruct + [>(proj1 … (Hind i4) H) // |>(proj2 … (Hind i4) (refl …)) //] ] qed. -lemma memb_append_l2: ∀S,a,l1,l2. - memb S a l2= true → memb S a (l1@l2) = true. -#S #a #l1 elim l1 normalize // -#b #tl #Hind #l2 cases (b==a) normalize /2/ -qed. - -let rec uniqueb (S:Alpha) l on l : bool ≝ - match l with - [ nil ⇒ true - | cons a tl ⇒ notb (memb S a tl) ∧ uniqueb S tl - ]. +definition DeqItem ≝ λS. + mk_DeqSet (pitem S) (beqitem S) (beqitem_true S). + +unification hint 0 ≔ S; + X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S) +(* ---------------------------------------- *) ⊢ + pitem S ≡ carr X. + +unification hint 0 ≔ S,i1,i2; + X ≟ mk_DeqSet (pitem S) (beqitem S) (beqitem_true S) +(* ---------------------------------------- *) ⊢ + beqitem S i1 i2 ≡ eqb X i1 i2. -definition sons ≝ λp:space Bin. - [〈move Bin true (\fst (\fst p)), move Bin true (\fst (\snd p))〉; - 〈move Bin false (\fst (\fst p)), move Bin false (\fst (\snd p))〉 - ]. +definition sons ≝ λS:DeqSet.λl:list S.λp:(pre S)×(pre S). + map ?? (λa.〈move S a (\fst (\fst p)),move S a (\fst (\snd p))〉) l. -axiom memb_sons: ∀p,q. memb (space Bin) p (sons q) = true → +lemma memb_sons: ∀S,l.∀p,q:(pre S)×(pre S). memb ? p (sons ? l q) = true → ∃a.(move ? a (\fst (\fst q)) = \fst p ∧ move ? a (\fst (\snd q)) = \snd p). - -(* -let rec test_sons (l:list (space Bin)) ≝ - match l with - [ nil ⇒ true - | cons hd tl ⇒ - beqb (\snd (\fst hd)) (\snd (\snd hd)) ∧ test_sons tl - ]. *) - -let rec unique_append (S:Alpha) (l1,l2: list S) on l1 ≝ - match l1 with - [ nil ⇒ l2 - | cons a tl ⇒ - let r ≝ unique_append S tl l2 in - if (memb S a r) then r else a::r - ]. - -lemma unique_append_unique: ∀S,l1,l2. uniqueb S l2 = true → - uniqueb S (unique_append S l1 l2) = true. -#S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2 -cases (true_or_false … (memb S a (unique_append S tl l2))) -#H >H normalize [@Hind //] >H normalize @Hind // +#S #l elim l [#p #q normalize in ⊢ (%→?); #abs @False_ind /2/] +#a #tl #Hind #p #q #H cases (orb_true_l … H) -H + [#H @(ex_intro … a) <(proj1 … (eqb_true …)H) /2/ + |#H @Hind @H + ] qed. -let rec bisim (n:nat) (frontier,visited: list (space Bin)) ≝ +let rec bisim S l n (frontier,visited: list ?) on n ≝ match n with [ O ⇒ 〈false,visited〉 (* assert false *) | S m ⇒ @@ -300,14 +227,14 @@ let rec bisim (n:nat) (frontier,visited: list (space Bin)) ≝ [ nil ⇒ 〈true,visited〉 | cons hd tl ⇒ if beqb (\snd (\fst hd)) (\snd (\snd hd)) then - bisim m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) - (sons hd)) tl) (hd::visited) + bisim S l m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) + (sons S l hd)) tl) (hd::visited) else 〈false,visited〉 ] ]. -lemma unfold_bisim: ∀n.∀frontier,visited: list (space Bin). - bisim n frontier visited = +lemma unfold_bisim: ∀S,l,n.∀frontier,visited: list ?. + bisim S l n frontier visited = match n with [ O ⇒ 〈false,visited〉 (* assert false *) | S m ⇒ @@ -315,85 +242,66 @@ lemma unfold_bisim: ∀n.∀frontier,visited: list (space Bin). [ nil ⇒ 〈true,visited〉 | cons hd tl ⇒ if beqb (\snd (\fst hd)) (\snd (\snd hd)) then - bisim m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) (sons hd)) tl) (hd::visited) + bisim S l m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) + (sons S l hd)) tl) (hd::visited) else 〈false,visited〉 ] ]. -#n cases n // qed. +#S #l #n cases n // qed. -lemma bisim_never: ∀frontier,visited: list (space Bin). - bisim O frontier visited = 〈false,visited〉. +lemma bisim_never: ∀S,l.∀frontier,visited: list ?. + bisim S l O frontier visited = 〈false,visited〉. #frontier #visited >unfold_bisim // qed. -lemma bisim_end: ∀m.∀visited: list (space Bin). - bisim (S m) [] visited = 〈true,visited〉. +lemma bisim_end: ∀Sig,l,m.∀visited: list ?. + bisim Sig l (S m) [] visited = 〈true,visited〉. #n #visisted >unfold_bisim // qed. -lemma bisim_step_true: ∀m.∀p.∀frontier,visited: list (space Bin). +lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: list ?. beqb (\snd (\fst p)) (\snd (\snd p)) = true → - bisim (S m) (p::frontier) visited = - bisim m (unique_append ? (filter ? (λx.notb(memb (space Bin) x (p::visited))) (sons p)) frontier) (p::visited). -#m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // + bisim Sig l (S m) (p::frontier) visited = + bisim Sig l m (unique_append ? (filter ? (λx.notb(memb ? x (p::visited))) + (sons Sig l p)) frontier) (p::visited). +#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // qed. -lemma bisim_step_false: ∀m.∀p.∀frontier,visited: list (space Bin). +lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: list ?. beqb (\snd (\fst p)) (\snd (\snd p)) = false → - bisim (S m) (p::frontier) visited = 〈false,visited〉. -#m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // + bisim Sig l (S m) (p::frontier) visited = 〈false,visited〉. +#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // qed. -definition visited_inv ≝ λe1,e2:pre Bin.λvisited: list (space Bin). +definition visited_inv ≝ λS.λe1,e2:pre S.λvisited: list ?. uniqueb ? visited = true ∧ ∀p. memb ? p visited = true → - (∃w.(moves Bin w e1 = \fst p) ∧ (moves Bin w e2 = \snd p)) ∧ + (∃w.(moves S w e1 = \fst p) ∧ (moves S w e2 = \snd p)) ∧ (beqb (\snd (\fst p)) (\snd (\snd p)) = true). -definition frontier_inv ≝ λfrontier,visited: list (space Bin). +definition frontier_inv ≝ λS.λfrontier,visited. uniqueb ? frontier = true ∧ -∀p. memb ? p frontier = true → +∀p:(pre S)×(pre S). memb ? p frontier = true → memb ? p visited = false ∧ ∃p1.((memb ? p1 visited = true) ∧ (∃a. move ? a (\fst (\fst p1)) = \fst p ∧ move ? a (\fst (\snd p1)) = \snd p)). -definition orb_true_r1: ∀b1,b2:bool. - b1 = true → (b1 ∨ b2) = true. -#b1 #b2 #H >H // qed. - -definition orb_true_r2: ∀b1,b2:bool. - b2 = true → (b1 ∨ b2) = true. -#b1 #b2 #H >H cases b1 // qed. - -definition orb_true_l: ∀b1,b2:bool. - (b1 ∨ b2) = true → (b1 = true) ∨ (b2 = true). -* normalize /2/ qed. - -definition andb_true_l1: ∀b1,b2:bool. - (b1 ∧ b2) = true → (b1 = true). -#b1 #b2 cases b1 normalize //. -qed. - -definition andb_true_l2: ∀b1,b2:bool. - (b1 ∧ b2) = true → (b2 = true). -#b1 #b2 cases b1 cases b2 normalize //. -qed. - -definition andb_true_l: ∀b1,b2:bool. +(* lemma andb_true: ∀b1,b2:bool. (b1 ∧ b2) = true → (b1 = true) ∧ (b2 = true). #b1 #b2 cases b1 normalize #H [>H /2/ |@False_ind /2/]. qed. -definition andb_true_r: ∀b1,b2:bool. +lemma andb_true_r: ∀b1,b2:bool. (b1 = true) ∧ (b2 = true) → (b1 ∧ b2) = true. #b1 #b2 cases b1 normalize * // -qed. +qed. *) lemma notb_eq_true_l: ∀b. notb b = true → b = false. #b cases b normalize // qed. +(* lemma notb_eq_true_r: ∀b. b = false → notb b = true. #b cases b normalize // qed. @@ -404,26 +312,9 @@ qed. lemma notb_eq_false_r:∀b. b = true → notb b = false. #b cases b normalize // -qed. - +qed. *) -axiom filter_true: ∀S,f,a,l. - memb S a (filter S f l) = true → f a = true. -(* -#S #f #a #l elim l [normalize #H @False_ind /2/] -#b #tl #Hind normalize cases (f b) normalize *) - -axiom memb_filter_memb: ∀S,f,a,l. - memb S a (filter S f l) = true → memb S a l = true. - -axiom unique_append_elim: ∀S:Alpha.∀P: S → Prop.∀l1,l2. -(∀x. memb S x l1 = true → P x) → (∀x. memb S x l2 = true → P x) → -∀x. memb S x (unique_append S l1 l2) = true → P x. - -axiom not_memb_to_not_eq: ∀S,a,b,l. - memb S a l = false → memb S b l = true → a==b = false. - -include "arithmetics/exp.ma". +(* include "arithmetics/exp.ma". *) let rec pos S (i:re S) on i ≝ match i with @@ -435,57 +326,6 @@ let rec pos S (i:re S) on i ≝ | k i ⇒ pos S i ]. -definition sublist ≝ - λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true. - -lemma memb_exists: ∀S,a,l.memb S a l = true → - ∃l1,l2.l=l1@(a::l2). -#S #a #l elim l [normalize #abs @False_ind /2/] -#b #tl #Hind #H cases (orb_true_l … H) - [#eqba @(ex_intro … (nil S)) @(ex_intro … tl) - >(proj1 … (eqb_true …) eqba) // - |#mem_tl cases (Hind mem_tl) #l1 * #l2 #eqtl - @(ex_intro … (b::l1)) @(ex_intro … l2) >eqtl // - ] -qed. - -lemma length_append: ∀A.∀l1,l2:list A. - |l1@l2| = |l1|+|l2|. -#A #l1 elim l1 // normalize /2/ -qed. - -lemma sublist_length: ∀S,l1,l2. - uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|. -#S #l1 elim l1 // -#a #tl #Hind #l2 #unique #sub -cut (∃l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //] -* #l3 * #l4 #eql2 >eql2 >length_append normalize -applyS le_S_S eql2 in sub; #sub #x #membx -cases (memb_append … (sub x (orb_true_r2 … membx))) - [#membxl3 @memb_append_l1 // - |#membxal4 cases (orb_true_l … membxal4) - [#eqax @False_ind cases (andb_true_l … unique) - >(proj1 … (eqb_true …) eqax) >membx normalize /2/ - |#membxl4 @memb_append_l2 // - ] - ] -qed. - -axiom memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true → -memb S x l = true ∧ (f x = true). - -axiom memb_filter_l: ∀S,f,l,x. memb S x l = true → (f x = true) → -memb S x (filter ? f l) = true. - -axiom sublist_unique_append_l1: - ∀S,l1,l2. sublist S l1 (unique_append S l1 l2). - -axiom sublist_unique_append_l2: - ∀S,l1,l2. sublist S l2 (unique_append S l1 l2). - -definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2. - foldr ?? (λi,acc.(map ?? (f i) l2)@acc) [ ] l1. let rec pitem_enum S (i:re S) on i ≝ match i with @@ -496,130 +336,146 @@ let rec pitem_enum S (i:re S) on i ≝ | c i1 i2 ⇒ compose ??? (pc S) (pitem_enum S i1) (pitem_enum S i2) | k i ⇒ map ?? (pk S) (pitem_enum S i) ]. - -axiom memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2. - memb S1 a1 l1 = true → memb S2 a2 l2 = true → - memb S3 (op a1 a2) (compose S1 S2 S3 op l1 l2) = true. -(* #S #op #a1 #a2 #l1 elim l1 [normalize //] -#x #tl #Hind #l2 elim l2 [#_ normalize #abs @False_ind /2/] -#y #tl2 #Hind2 #membx #memby normalize *) - -axiom pitem_enum_complete: ∀i: pitem Bin. - memb BinItem i (pitem_enum ? (forget ? i)) = true. -(* -#i elim i - [// - |// - |* // - |* // - |#i1 #i2 #Hind1 #Hind2 @memb_compose // - |#i1 #i2 #Hind1 #Hind2 @memb_compose // - | -*) + +lemma pitem_enum_complete : ∀S.∀i:pitem S. + memb (DeqItem S) i (pitem_enum S (|i|)) = true. +#S #i elim i + [1,2:// + |3,4:#c normalize >(\b (refl … c)) // + |5,6:#i1 #i2 #Hind1 #Hind2 @(memb_compose (DeqItem S) (DeqItem S)) // + |#i #Hind @(memb_map (DeqItem S)) // + ] +qed. definition pre_enum ≝ λS.λi:re S. compose ??? (λi,b.〈i,b〉) (pitem_enum S i) [true;false]. + +lemma pre_enum_complete : ∀S.∀e:pre S. + memb ? e (pre_enum S (|\fst e|)) = true. +#S * #i #b @(memb_compose (DeqItem S) DeqBool ? (λi,b.〈i,b〉)) +// cases b normalize // +qed. definition space_enum ≝ λS.λi1,i2:re S. - compose ??? (λe1,e2.〈e1,e2〉) (pre_enum S i1) (pre_enum S i1). + compose ??? (λe1,e2.〈e1,e2〉) (pre_enum S i1) (pre_enum S i2). + +lemma space_enum_complete : ∀S.∀e1,e2: pre S. + memb ? 〈e1,e2〉 (space_enum S (|\fst e1|) (|\fst e2|)) = true. +#S #e1 #e2 @(memb_compose … (λi,b.〈i,b〉)) +// qed. -axiom space_enum_complete : ∀S.∀e1,e2: pre S. - memb (space S) 〈e1,e2〉 (space_enum S (|\fst e1|) (|\fst e2|)) = true. +definition visited_inv_1 ≝ λS.λe1,e2:pre S.λvisited: list ?. +uniqueb ? visited = true ∧ + ∀p. memb ? p visited = true → + ∃w.(moves S w e1 = \fst p) ∧ (moves S w e2 = \snd p). -lemma bisim_ok1: ∀e1,e2:pre Bin.\sem{e1}=1\sem{e2} → - ∀n.∀frontier,visited:list (space Bin). - |space_enum Bin (|\fst e1|) (|\fst e2|)| < n + |visited|→ - visited_inv e1 e2 visited → frontier_inv frontier visited → - \fst (bisim n frontier visited) = true. -#e1 #e2 #same #n elim n +lemma bisim_ok1: ∀S.∀e1,e2:pre S.\sem{e1}=1\sem{e2} → + ∀l,n.∀frontier,visited:list (*(space S) *) ((pre S)×(pre S)). + |space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|→ + visited_inv_1 S e1 e2 visited → frontier_inv S frontier visited → + \fst (bisim S l n frontier visited) = true. +#Sig #e1 #e2 #same #l #n elim n [#frontier #visited #abs * #unique #H @False_ind @(absurd … abs) @le_to_not_lt @sublist_length // * #e11 #e21 #membp cut ((|\fst e11| = |\fst e1|) ∧ (|\fst e21| = |\fst e2|)) [|* #H1 #H2

fst_eq >snd_eq #we1 #we2 #_ - same_kernel_moves // |#m #HI * [#visited #vinv #finv >bisim_end //] #p #front_tl #visited #Hn * #u_visited #vinv * #u_frontier #finv cases (finv p (memb_hd …)) #Hp * #p2 * #visited_p2 - * #a * #movea1 #movea2 - cut (∃w.(moves Bin w e1 = \fst p) ∧ (moves Bin w e2 = \snd p)) - [cases (vinv … visited_p2) -vinv * #w1 * #mw1 #mw2 #_ - @(ex_intro … (w1@[a])) /2/] + * #a * #movea1 #movea2 + cut (∃w.(moves Sig w e1 = \fst p) ∧ (moves Sig w e2 = \snd p)) + [cases (vinv … visited_p2) -vinv #w1 * #mw1 #mw2 + @(ex_intro … (w1@[a])) % //] -movea2 -movea1 -a -visited_p2 -p2 #reachp cut (beqb (\snd (\fst p)) (\snd (\snd p)) = true) [cases reachp #w * #move_e1 #move_e2 (bisim_step_true … ptest) @HI -HI - [(bisim_step_true … ptest) @HI -HI + [Hp whd in ⊢ (??%?); //] + #p1 #H (cases (orb_true_l … H)) + [#eqp <(\P eqp) // |#visited_p1 @(vinv … visited_p1) ] - |whd % [@unique_append_unique @(andb_true_l2 … u_frontier)] - @unique_append_elim #q #H - [% - [@notb_eq_true_l @(filter_true … H) - |@(ex_intro … p) % // - @(memb_sons … (memb_filter_memb … H)) + |whd % [@unique_append_unique @(andb_true_r … u_frontier)] + @unique_append_elim #q #H + [% + [@notb_eq_true_l @(filter_true … H) + |@(ex_intro … p) % [@memb_hd|@(memb_sons … (memb_filter_memb … H))] + ] + |cases (finv q ?) [|@memb_cons //] + #nvq * #p1 * #Hp1 #reach % + [cut ((p==q) = false) [|#Hpq whd in ⊢ (??%?); >Hpq @nvq] + cases (andb_true … u_frontier) #notp #_ + @(not_memb_to_not_eq … H) @notb_eq_true_l @notp + |cases (proj2 … (finv q ?)) + [#p1 * #Hp1 #reach @(ex_intro … p1) % // @memb_cons // + |@memb_cons // ] - |cases (finv q ?) [|@memb_cons //] - #nvq * #p1 * #Hp1 #reach % - [cut ((p==q) = false) [|#Hpq whd in ⊢ (??%?); >Hpq @nvq] - cases (andb_true_l … u_frontier) #notp #_ - @(not_memb_to_not_eq … H) @notb_eq_true_l @notp - |cases (proj2 … (finv q ?)) - [#p1 * #Hp1 #reach @(ex_intro … p1) % // @memb_cons // - |@memb_cons // - ] - ] - ] - ] + ] + ] ] ] qed. -definition all_true ≝ λl.∀p. memb (space Bin) p l = true → +definition all_true ≝ λS.λl.∀p:(pre S) × (pre S). memb ? p l = true → (beqb (\snd (\fst p)) (\snd (\snd p)) = true). -definition sub_sons ≝ λl1,l2.∀x,a. -memb (space Bin) x l1 = true → - memb (space Bin) 〈move ? a (\fst (\fst x)), move ? a (\fst (\snd x))〉 l2 = true. +definition sub_sons ≝ λS,l,l1,l2.∀x:(pre S) × (pre S).∀a:S. +memb ? x l1 = true → memb S a l = true → + memb ? 〈move ? a (\fst (\fst x)), move ? a (\fst (\snd x))〉 l2 = true. lemma reachable_bisim: - ∀n.∀frontier,visited,visited_res:list (space Bin). - all_true visited → - sub_sons visited (frontier@visited) → - bisim n frontier visited = 〈true,visited_res〉 → - (sub_sons visited_res visited_res ∧ + ∀S,l,n.∀frontier,visited,visited_res:list ?. + all_true S visited → + sub_sons S l visited (frontier@visited) → + bisim S l n frontier visited = 〈true,visited_res〉 → + (sub_sons S l visited_res visited_res ∧ sublist ? visited visited_res ∧ - all_true visited_res). -#n elim n + all_true S visited_res). +#S #l #n elim n [#fron #vis #vis_res #_ #_ >bisim_never #H destruct |#m #Hind * - [-Hind #vis #vis_res #allv #H normalize in ⊢ (%→?); - #H1 destruct % // % // #p /2/ + [(* case empty frontier *) + -Hind #vis #vis_res #allv #H normalize in ⊢ (%→?); + #H1 destruct % // % // #p /2 by / |#hd cases (true_or_false (beqb (\snd (\fst hd)) (\snd (\snd hd)))) - [|#H #tl #vis #vis_res #allv >(bisim_step_false … H) #_ #H1 destruct] + [|(* case head of the frontier is non ok (absurd) *) + #H #tl #vis #vis_res #allv >(bisim_step_false … H) #_ #H1 destruct] + (* frontier = hd:: tl and hd is ok *) #H #tl #visited #visited_res #allv >(bisim_step_true … H) - cut (all_true (hd::visited)) - [#p #H cases (orb_true_l … H) - [#eqp <(proj1 … (eqb_true …) eqp) // |@allv]] + (* new_visited = hd::visited are all ok *) + cut (all_true S (hd::visited)) + [#p #H1 cases (orb_true_l … H1) [#eqp <(\P eqp) @H |@allv]] + (* we now exploit the induction hypothesis *) #allh #subH #bisim cases (Hind … allh … bisim) -Hind - [* #H1 #H2 #H3 % // % // #p #H4 @H2 @memb_cons //] - #x #a #membx + [* #H1 #H2 #H3 % // % // #p #H4 @H2 @memb_cons //] + (* the only thing left to prove is the sub_sons invariant *) + #x #a #membx #memba cases (orb_true_l … membx) - [#eqhdx >(proj1 … (eqb_true …) eqhdx) - letin xa ≝ 〈move Bin a (\fst (\fst x)), move Bin a (\fst (\snd x))〉 - cases (true_or_false … (memb (space Bin) xa (x::visited))) - [#membxa @memb_append_l2 // - |#membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l - [whd in ⊢ (??(??%%)?); cases a [@memb_hd |@memb_cons @memb_hd] - |>membxa // + [(* case x = hd *) + #eqhdx >(proj1 … (eqb_true …) eqhdx) + (* xa is the son of x w.r.t. a; we must distinguish the case xa + was already visited form the case xa is new *) + letin xa ≝ 〈move S a (\fst (\fst x)), move S a (\fst (\snd x))〉 + cases (true_or_false … (memb ? xa (x::visited))) + [(* xa visited - trivial *) #membxa @memb_append_l2 // + |(* xa new *) #membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l + [>membxa // + |(* this can be probably improved *) + generalize in match memba; -memba elim l + [whd in ⊢ (??%?→?); #abs @False_ind /2/ + |#b #others #Hind #memba cases (orb_true_l … memba) #H + [>(proj1 … (eqb_true …) H) @memb_hd + |@memb_cons @Hind // + ] + ] ] ] - |#H1 letin xa ≝ 〈move Bin a (\fst (\fst x)), move Bin a (\fst (\snd x))〉 - cases (memb_append … (subH x a H1)) + |(* case x in visited *) + #H1 letin xa ≝ 〈move S a (\fst (\fst x)), move S a (\fst (\snd x))〉 + cases (memb_append … (subH x a H1 memba)) [#H2 (cases (orb_true_l … H2)) [#H3 @memb_append_l2 >(proj1 … (eqb_true …) H3) @memb_hd |#H3 @memb_append_l1 @sublist_unique_append_l2 @H3 @@ -629,35 +485,155 @@ lemma reachable_bisim: ] ] ] -qed. +qed. + +(* pit state *) +let rec blank_item (S: DeqSet) (i: re S) on i :pitem S ≝ + match i with + [ z ⇒ `∅ + | e ⇒ ϵ + | s y ⇒ `y + | o e1 e2 ⇒ (blank_item S e1) + (blank_item S e2) + | c e1 e2 ⇒ (blank_item S e1) · (blank_item S e2) + | k e ⇒ (blank_item S e)^* ]. + +definition pit_pre ≝ λS.λi.〈blank_item S (|i|), false〉. + +let rec occur (S: DeqSet) (i: re S) on i ≝ + match i with + [ z ⇒ [ ] + | e ⇒ [ ] + | s y ⇒ [y] + | o e1 e2 ⇒ unique_append ? (occur S e1) (occur S e2) + | c e1 e2 ⇒ unique_append ? (occur S e1) (occur S e2) + | k e ⇒ occur S e]. + +axiom memb_single: ∀S,a,x. memb S a [x] = true → a = x. + +axiom tech: ∀b. b ≠ true → b = false. +axiom tech2: ∀b. b = false → b ≠ true. + +lemma not_occur_to_pit: ∀S,a.∀i:pitem S. memb S a (occur S (|i|)) = false → + move S a i = pit_pre S i. +#S #a #i elim i // + [#x cases (true_or_false (a==x)) + [#H >(proj1 …(eqb_true …) H) whd in ⊢ ((??%?)→?); + >(proj2 …(eqb_true …) (refl …)) whd in ⊢ ((??%?)→?); #abs @False_ind /2/ + |#H normalize >H // + ] + |#i1 #i2 #Hind1 #Hind2 #H >move_cat >Hind1 [2:@tech + @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l1 //] + >Hind2 [2:@tech @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l2 //] + // + |#i1 #i2 #Hind1 #Hind2 #H >move_plus >Hind1 [2:@tech + @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l1 //] + >Hind2 [2:@tech @(not_to_not … (tech2 … H)) #H1 @sublist_unique_append_l2 //] + // + |#i #Hind #H >move_star >Hind // @H + ] +qed. + +lemma move_pit: ∀S,a,i. move S a (\fst (pit_pre S i)) = pit_pre S i. +#S #a #i elim i // + [#i1 #i2 #Hind1 #Hind2 >move_cat >Hind1 >Hind2 // + |#i1 #i2 #Hind1 #Hind2 >move_plus >Hind1 >Hind2 // + |#i #Hind >move_star >Hind // + ] +qed. -axiom bisim_char: ∀e1,e2:pre Bin. -(∀w.(beqb (\snd (moves ? w e1)) (\snd (moves ? w e2))) = true) → +lemma moves_pit: ∀S,w,i. moves S w (pit_pre S i) = pit_pre S i. +#S #w #i elim w // #a #tl >moves_cons // +qed. + +lemma to_pit: ∀S,w,e. ¬ sublist S w (occur S (|\fst e|)) → + moves S w e = pit_pre S (\fst e). +#S #w elim w + [(* orribile *) + #e * #H @False_ind @H normalize #a #abs @False_ind /2/ + |#a #tl #Hind #e #H cases (true_or_false (memb S a (occur S (|\fst e|)))) + [#Htrue >moves_cons whd in ⊢ (???%); <(same_kernel … a) + @Hind >same_kernel @(not_to_not … H) #H1 #b #memb cases (orb_true_l … memb) + [#H2 <(proj1 … (eqb_true …) H2) // |#H2 @H1 //] + |#Hfalse >moves_cons >not_occur_to_pit // + ] + ] +qed. + +definition occ ≝ λS.λe1,e2:pre S. + unique_append ? (occur S (|\fst e1|)) (occur S (|\fst e2|)). + +(* definition occS ≝ λS:DeqSet.λoccur. + PSig S (λx.memb S x occur = true). *) + +lemma occ_enough: ∀S.∀e1,e2:pre S. +(∀w.(sublist S w (occ S e1 e2))→ + (beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true) \to +∀w.(beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true. +#S #e1 #e2 #H #w +cut (sublist S w (occ S e1 e2) ∨ ¬(sublist S w (occ S e1 e2))) +[elim w + [%1 #a normalize in ⊢ (%→?); #abs @False_ind /2/ + |#a #tl * #subtl + [cases (true_or_false (memb S a (occ S e1 e2))) #memba + [%1 whd #x #membx cases (orb_true_l … membx) + [#eqax <(proj1 … (eqb_true …) eqax) // + |@subtl + ] + |%2 @(not_to_not … (tech2 … memba)) #H1 @H1 @memb_hd + ] + |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons // + ] + ] +|* [@H] + #H >to_pit + [2: @(not_to_not … H) #H1 #a #memba @sublist_unique_append_l1 @H1 //] + >to_pit + [2: @(not_to_not … H) #H1 #a #memba @sublist_unique_append_l2 @H1 //] + // +] +qed. + +lemma bisim_char: ∀S.∀e1,e2:pre S. +(∀w.(beqb (\snd (moves S w e1)) (\snd (moves ? w e2))) = true) → \sem{e1}=1\sem{e2}. +#S #e1 #e2 #H @(proj2 … (equiv_sem …)) #w @(\P ?) @H +qed. -lemma bisim_ok2: ∀e1,e2:pre Bin. - (beqb (\snd e1) (\snd e2) = true) → - ∀n.∀frontier:list (space Bin). - sub_sons [〈e1,e2〉] (frontier@[〈e1,e2〉]) → - \fst (bisim n frontier [〈e1,e2〉]) = true → \sem{e1}=1\sem{e2}. -#e1 #e2 #Hnil #n #frontier #init #bisim_true -letin visited_res ≝ (\snd (bisim n frontier [〈e1,e2〉])) -cut (bisim n frontier [〈e1,e2〉] = 〈true,visited_res〉) +lemma bisim_ok2: ∀S.∀e1,e2:pre S. + (beqb (\snd e1) (\snd e2) = true) → ∀n. + \fst (bisim S (occ S e1 e2) n (sons S (occ S e1 e2) 〈e1,e2〉) [〈e1,e2〉]) = true → + \sem{e1}=1\sem{e2}. +#S #e1 #e2 #Hnil #n +letin rsig ≝ (occ S e1 e2) +letin frontier ≝ (sons S rsig 〈e1,e2〉) +letin visited_res ≝ (\snd (bisim S rsig n frontier [〈e1,e2〉])) +#bisim_true +cut (bisim S rsig n frontier [〈e1,e2〉] = 〈true,visited_res〉) [fst_eq >snd_eq >moves_cons >moves_cons - @(Hind 〈?,?〉) @(H1 〈?,?〉) //] #all_reach -@bisim_char #w @(H3 〈?,?〉) @(all_reach w 〈?,?〉) @H2 // +cut (∀w.sublist ? w (occ S e1 e2)→∀p.memb ? p visited_res = true → + memb ? 〈moves ? w (\fst p), moves ? w (\snd p)〉 visited_res = true) + [#w elim w [#_ #p #H4 >moves_empty >moves_empty moves_cons >moves_cons + @(Hind ? 〈?,?〉) [#x #H4 @Hsub @memb_cons //] + @(H1 〈?,?〉) [@visp| @Hsub @memb_hd]] #all_reach +@bisim_char @occ_enough +#w #Hsub @(H3 〈?,?〉) @(all_reach w Hsub 〈?,?〉) @H2 // qed. +(* definition tt ≝ ps Bin true. definition ff ≝ ps Bin false. definition eps ≝ pe Bin. @@ -667,10 +643,8 @@ definition exp2 ≝ ff · (eps + tt). definition exp3 ≝ move Bin true (\fst (•exp1)). definition exp4 ≝ move Bin true (\fst (•exp2)). definition exp5 ≝ move Bin false (\fst (•exp1)). -definition exp6 ≝ move Bin false (\fst (•exp2)). +definition exp6 ≝ move Bin false (\fst (•exp2)). *) + -example comp1 : bequiv 15 (•exp1) (•exp2) [ ] = false . -normalize // -qed.