From: matitaweb Date: Tue, 6 Mar 2012 18:01:27 +0000 (+0000) Subject: commit by user andrea X-Git-Tag: make_still_working~1894 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=01ce59a8a53aad0e7375fcc76125842382480a59 commit by user andrea --- diff --git a/weblib/tutorial/chapter10.ma b/weblib/tutorial/chapter10.ma index f52e80570..9eef6a350 100644 --- a/weblib/tutorial/chapter10.ma +++ b/weblib/tutorial/chapter10.ma @@ -1,57 +1,84 @@ -include "cahpter9.ma". +(* +h1Equivalence/h1*) -(* bisimulation *) -definition cofinal ≝ λS.λp:(pre S)×(pre S). - \snd (\fst p) = \snd (\snd p). - -theorem equiv_sem: ∀S:DeqSet.∀e1,e2:pre S. - \sem{e1} =1 \sem{e2} ↔ ∀w.cofinal ? 〈moves ? w e1,moves ? w e2〉. +include "tutorial/chapter9.ma". + +(* We say that two pres \langle i_1,b_1\rangle$ and +$\langle i_1,b_1\rangle$ are {\em cofinal} if and only if +$b_1 = b_2$. *) + +definition cofinal ≝ λS.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). + a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p). + +(* As a corollary of decidable_sem, we have that two expressions +e1 and e2 are equivalent iff for any word w the states reachable +through w are cofinal. *) + +theorem equiv_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2} a title="iff" href="cic:/fakeuri.def(1)"↔/a ∀w.a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉. #S #e1 #e2 % [#same_sem #w - cut (∀b1,b2. iff (b1 = true) (b2 = true) → (b1 = b2)) - [* * // * #H1 #H2 [@sym_eq @H1 //| @H2 //]] - #Hcut @Hcut @iff_trans [|@decidable_sem] - @iff_trans [|@same_sem] @iff_sym @decidable_sem -|#H #w1 @iff_trans [||@decidable_sem] 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 //] +cases (a href="cic:/matita/tutorial/chapter5/decidable_sublist.def(6)"decidable_sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2)) [@H] -H #H + >a href="cic:/matita/tutorial/chapter9/to_pit.def(10)"to_pit/a [2: @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 #a #memba @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a @H1 //] + >a href="cic:/matita/tutorial/chapter9/to_pit.def(10)"to_pit/a [2: @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 #a #memba @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"sublist_unique_append_l2/a @H1 //] // qed. -lemma equiv_sem_occ: ∀S.∀e1,e2:pre S. -(∀w.(sublist S w (occ S e1 e2))→ cofinal ? 〈moves ? w e1,moves ? w e2〉) -→ \sem{e1}=1\sem{e2}. -#S #e1 #e2 #H @(proj2 … (equiv_sem …)) @occ_enough #w @H +(* The following is a stronger version of equiv_sem, relative to characters +occurring the given regular expressions. *) + +lemma equiv_sem_occ: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +(∀w.(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2))→ a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉) +→ a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +#S #e1 #e2 #H @(a href="cic:/matita/basics/logic/proj2.def(2)"proj2/a … (a href="cic:/matita/tutorial/chapter10/equiv_sem.def(16)"equiv_sem/a …)) @a href="cic:/matita/tutorial/chapter10/occ_enough.def(11)"occ_enough/a #w @H qed. -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. +(* +h2Bisimulations/h2 + +We say that a list of pairs of pres is a bisimulation if it is closed +w.r.t. moves, and all its members are cofinal. +*) + +definition sons ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). + a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (λa.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)),a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p))〉) l. -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). -#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) >(\P H) /2/ |#H @Hind @H] +lemma memb_sons: ∀S,l.∀p,q:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? l q) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → + a title="exists" href="cic:/fakeuri.def(1)"∃/aa.(a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a q)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p a title="logical and" href="cic:/fakeuri.def(1)"∧/a + a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a q)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p). +#S #l elim l [#p #q normalize in ⊢ (%→?); #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] +#a #tl #Hind #p #q #H cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … H) -H + [#H @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … a) >(\P H) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ |#H @Hind @H] qed. -definition is_bisim ≝ λS:DeqSet.λl:list ?.λalpha:list S. - ∀p:(pre S)×(pre S). memb ? p l = true → cofinal ? p ∧ (sublist ? (sons ? alpha p) l). +definition is_bisim ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.λalpha:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S. + ∀p:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? p a title="logical and" href="cic:/fakeuri.def(1)"∧/a (a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a ? (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? alpha p) l). -lemma bisim_to_sem: ∀S:DeqSet.∀l:list ?.∀e1,e2: pre S. - is_bisim S l (occ S e1 e2) → memb ? 〈e1,e2〉 l = true → \sem{e1}=1\sem{e2}. -#S #l #e1 #e2 #Hbisim #Hmemb @equiv_sem_occ -#w #Hsub @(proj1 … (Hbisim 〈moves S w e1,moves S w e2〉 ?)) +(* Using lemma equiv_sem_occ it is easy to prove the following result: *) + +lemma bisim_to_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"is_bisim/a S l (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2〉 l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +#S #l #e1 #e2 #Hbisim #Hmemb @a href="cic:/matita/tutorial/chapter10/equiv_sem_occ.def(17)"equiv_sem_occ/a +#w #Hsub @(a href="cic:/matita/basics/logic/proj1.def(2)"proj1/a … (Hbisim a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e2〉 ?)) lapply Hsub @(list_elim_left … w) [//] #a #w1 #Hind #Hsub >moves_left >moves_left @(proj2 …(Hbisim …(Hind ?))) [#x #Hx @Hsub @memb_append_l1 // @@ -60,7 +87,31 @@ lapply Hsub @(list_elim_left … w) [//] ] qed. -(* the algorithm *) +(* This is already an interesting result: checking if l is a bisimulation +is decidable, hence we could generate l with some untrusted piece of code +and then run a (boolean version of) is_bisim to check that it is actually +a bisimulation. +However, in order to prove that equivalence of regular expressions +is decidable we must prove that we can always effectively build such a list +(or find a counterexample). +The idea is that the list we are interested in is just the set of +all pair of pres reachable from the initial pair via some +sequence of moves. + +The algorithm for computing reachable nodes in graph is a very +traditional one. We split nodes in two disjoint lists: a list of +visited nodes and a frontier, composed by all nodes connected +to a node in visited but not visited already. At each step we select a node +a from the frontier, compute its sons, add a to the set of +visited nodes and the (not already visited) sons to the frontier. + +Instead of fist computing reachable nodes and then performing the +bisimilarity test we can directly integrate it in the algorithm: +the set of visited nodes is closed by construction w.r.t. reachability, +so we have just to check cofinality for any node we add to visited. + +Here is the extremely simple algorithm: *) + let rec bisim S l n (frontier,visited: list ?) on n ≝ match n with [ O ⇒ 〈false,visited〉 (* assert false *) @@ -74,7 +125,15 @@ let rec bisim S l n (frontier,visited: list ?) on n ≝ else 〈false,visited〉 ] ]. - + +(* The integer n is an upper bound to the number of recursive call, +equal to the dimension of the graph. It returns a pair composed +by a boolean and a the set of visited nodes; the boolean is true +if and only if all visited nodes are cofinal. + +The following results explicitly state the behaviour of bisim is the general +case and in some relevant instances *) + lemma unfold_bisim: ∀S,l,n.∀frontier,visited: list ?. bisim S l n frontier visited = match n with @@ -90,7 +149,7 @@ lemma unfold_bisim: ∀S,l,n.∀frontier,visited: list ?. ] ]. #S #l #n cases n // qed. - + lemma bisim_never: ∀S,l.∀frontier,visited: list ?. bisim S l O frontier visited = 〈false,visited〉. #frontier #visited >unfold_bisim // @@ -115,9 +174,12 @@ beqb (\snd (\fst p)) (\snd (\snd p)) = false → #Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // qed. -lemma notb_eq_true_l: ∀b. notb b = true → b = false. +(* In order to prove termination of bisim we must be able to effectively +enumerate all possible pres: *) + +(* lemma notb_eq_true_l: ∀b. notb b = true → b = false. #b cases b normalize // -qed. +qed. *) let rec pitem_enum S (i:re S) on i ≝ match i with @@ -164,6 +226,11 @@ uniqueb ? l = true ∧ definition disjoint ≝ λS:DeqSet.λl1,l2. ∀p:S. memb S p l1 = true → memb S p l2 = false. +(* We are ready to prove that bisim is correct; we use the invariant +that at each call of bisim the two lists visited and frontier only contain +nodes reachable from \langle e_1,e_2\rangle, hence it is absurd to suppose +to meet a pair which is not cofinal. *) + lemma bisim_correct: ∀S.∀e1,e2:pre S.\sem{e1}=1\sem{e2} → ∀l,n.∀frontier,visited:list ((pre S)×(pre S)). |space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|→ @@ -206,7 +273,11 @@ lemma bisim_correct: ∀S.∀e1,e2:pre S.\sem{e1}=1\sem{e2} → ] ] ] -qed. +qed. + +(* For completeness, we use the invariant that all the nodes in visited are cofinal, +and the sons of visited are either in visited or in the frontier; since +at the end frontier is empty, visited is hence a bisimulation. *) definition all_true ≝ λS.λl.∀p:(pre S) × (pre S). memb ? p l = true → (beqb (\snd (\fst p)) (\snd (\snd p)) = true). @@ -268,6 +339,10 @@ lemma bisim_complete: ] qed. +(* We can now give the definition of the equivalence algorithm, and +prove that two expressions are equivalente if and only if they define +the same language. *) + definition equiv ≝ λSig.λre1,re2:re Sig. let e1 ≝ •(blank ? re1) in let e2 ≝ •(blank ? re2) in @@ -322,5 +397,3 @@ normalize // qed. - - diff --git a/weblib/tutorial/chapter5.ma b/weblib/tutorial/chapter5.ma index a3b88978b..8d4fc9f6c 100644 --- a/weblib/tutorial/chapter5.ma +++ b/weblib/tutorial/chapter5.ma @@ -4,18 +4,18 @@ effectively searching an element of that set inside a data structure. In this section we shall define several boolean functions acting on lists of elements in a DeqSet, and prove some of their properties.*) -include "basics/list.ma". +include "basics/list.ma". include "tutorial/chapter4.ma". (* The first function we define is an effective version of the membership relation, between an element x and a list l. Its definition is a straightforward recursion on l.*) -let rec memb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (x:S) (l: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span S) on l ≝ +let rec memb (S:DeqSet) (x:S) (l: listspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span S) on l ≝ match l with - [ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a - | cons a tl ⇒ (x a title="eqb" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/span= a) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a memb S x tl - ]. + [ nil ⇒ false + | cons a tl ⇒ (x == a) ∨ memb S x tl + ]span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/spanspan class="error" title="No choices for ID nil"/span. notation < "\memb x l" non associative with precedence 90 for @{'memb $x $l}. interpretation "boolean membership" 'memb a l = (memb ? a l). @@ -35,71 +35,71 @@ interpretation "boolean membership" 'memb a l = (memb ? a l). (op a b) is a member of (compose op l1 l2) *) -lemma memb_hd: ∀S,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (aa title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #a #l normalize >(a href="cic:/matita/basics/logic/proj2.def(2)"proj2/a … (a href="cic:/matita/tutorial/chapter4/eqb_true.fix(0,0,4)"eqb_true/a S …) (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a S a)) // +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. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S a (ba title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #a #b #l normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize // + memb S a l = true → membspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S a (b::l) = true. +#S #a #b #l normalize cases (a==b) normalize // qed. -lemma memb_single: ∀S,a,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (xa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. -#S #a #x normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) #H - [#_ >(\P H) // |>H normalize #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] +lemma memb_single: ∀S,a,x. memb S a (x::[]) = true → a = x. +#S #a #x normalize cases (true_or_false … (a==x)) #H + [#_ >(\P H) // |>H normalize #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/] qed. lemma memb_append: ∀S,a,l1,l2. -a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/aspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/spanl2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #a #l1 elim l1 normalize [#l2 #H %2 //] -#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a/span/span/ +memb S a (l1@span class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/spanl2) = true → memb S a l1= true ∨ memb S a l2 = true. +#S #a #l1span class="error" title="Parse error: illegal begin of statement"/spanspan class="error" title="Parse error: illegal begin of statement"/span elim l1 normalize [#l2 #H %2 //] +#b #tl #Hind #l2 cases (a==b) normalize /span class="autotactic"2span class="autotrace" trace orb_true_l/span/span/ qed. lemma memb_append_l1: ∀S,a,l1,l2. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/al2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. + memb S a l1= true → memb S a (l1@l2) = true. #S #a #l1 elim l1 normalize - [normalize #le #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ - |#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ + [normalize #le #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/ + |#b #tl #Hind #l2 cases (a==b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ ] qed. lemma memb_append_l2: ∀S,a,l1,l2. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/al2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. + memb S a l2= true → memb S a (l1@l2) = true. #S #a #l1 elim l1 normalize // -#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ +#b #tl #Hind #l2 cases (a==b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -lemma memb_exists: ∀S,a,l.a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span → a title="exists" href="cic:/fakeuri.def(1)"∃/al1,l2.la title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al1a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:l2). -#S #a #l elim l [normalize #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] -#b #tl #Hind #H cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … H) - [#eqba @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a S)) @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … tl) >(\P eqba) // +lemma memb_exists: ∀S,a,l.memb S a l = truespan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span → ∃l1,l2.l=l1@(a::l2). +#S #a #l elim l [normalize #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/] +#b #tl #Hind #H cases (orb_true_l … H) + [#eqba @(ex_intro … (nil S)) @(ex_intro … tl) >(\P eqba) // |#mem_tl cases (Hind mem_tl) #l1 * #l2 #eqtl - @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (ba title="cons" href="cic:/fakeuri.def(1)":/a:l1)) @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … l2) >eqtl // + @(ex_intro … (b::l1)) @(ex_intro … l2) >eqtl // ] qed. lemma not_memb_to_not_eq: ∀S,a,b,l. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S b l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. -#S #a #b #l cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) // -#eqab >(\P eqab) #H >H #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ + memb S a l = false → memb S b l = true → a==b = false. +#S #a #b #l cases (true_or_false (a==b)) // +#eqab >(\P eqab) #H >H #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/ qed. -lemma memb_map: ∀S1,S2,f,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S1 a la title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S2 (f a) (a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a … f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +lemma memb_map: ∀S1,S2,f,a,l. memb S1 a l= true → + memb S2 (f a) (map … f l) = true. #S1 #S2 #f #a #l elim l normalize [//] -#x #tl #memba cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) - [#eqx >eqx >(\P eqx) >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … (f x))) normalize // - |#eqx >eqx cases (f aa title="eqb" href="cic:/fakeuri.def(1)"=/a=f x) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ +#x #tl #memba cases (true_or_false (a==x)) + [#eqx >eqx >(\P eqx) >(\b (refl … (f x))) normalize // + |#eqx >eqx cases (f a==f x) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ ] qed. lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S1 a1 l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S2 a2 l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S3 (op a1 a2) (a href="cic:/matita/basics/list/compose.def(2)"compose/a S1 S2 S3 op l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. + memb S1 a1 l1 = true → memb S2 a2 l2 = true → + memb S3 (op a1 a2) (compose S1 S2 S3 op l1 l2) = true. #S1 #S2 #S3 #op #a1 #a2 #l1 elim l1 [normalize //] -#x #tl #Hind #l2 #memba1 #memba2 cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … memba1) - [#eqa1 >(\P eqa1) @a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"memb_append_l1/a @a href="cic:/matita/tutorial/chapter5/memb_map.def(5)"memb_map/a // - |#membtl @a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"memb_append_l2/a @Hind // +#x #tl #Hind #l2 #memba1 #memba2 cases (orb_true_l … memba1) + [#eqa1 >(\P eqa1) @memb_append_l1 @memb_map // + |#membtl @memb_append_l2 @Hind // ] qed. @@ -108,84 +108,84 @@ to avoid duplications of elements. The following uniqueb check this property. *) (*************** unicity test *****************) -let rec uniqueb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) l on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +let rec uniqueb (S:DeqSet) l on l : bool ≝ match l with - [ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a - | cons a tl ⇒ a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a tl) a title="boolean and" href="cic:/fakeuri.def(1)"∧/a uniqueb S tl + [ nil ⇒ true + | cons a tl ⇒ notb (memb S a tl) ∧ uniqueb S tl ]. (* unique_append l1 l2 add l1 in fornt of l2, but preserving unicity *) -let rec unique_append (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l1,l2: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S) on l1 ≝ +let rec unique_append (S:DeqSet) (l1,l2: list S) on l1 ≝ match l1 with [ nil ⇒ l2 | cons a tl ⇒ let r ≝ unique_append S tl l2 in - if a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a r then r else aa title="cons" href="cic:/fakeuri.def(1)":/a:r + if memb S a r then r else a::r ]. -axiom unique_append_elim: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀P: S → Prop.∀l1,l2. -(∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/span a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x) → (∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x) → -∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x. +axiom unique_append_elim: ∀S:DeqSet.∀P: S → Prop.∀l1,l2. +(∀x. memb S x l1 =span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/span true → P x) → (∀x. memb S x l2 = true → P x) → +∀x. memb S x (unique_append S l1 l2) = true → P x. -lemma unique_append_unique: ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +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 (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2))) +cases (true_or_false … (memb S a (unique_append S tl l2))) #H >H normalize [@Hind //] >H normalize @Hind // qed. (******************* sublist *******************) definition sublist ≝ - λS,l1,l2.∀a. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. + λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true. lemma sublist_length: ∀S,l1,l2. - a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2 → a title="norm" href="cic:/fakeuri.def(1)"|/al1| a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a title="norm" href="cic:/fakeuri.def(1)"|/al2|. + uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|. #S #l1 elim l1 // #a #tl #Hind #l2 #unique #sub -cut (a title="exists" href="cic:/fakeuri.def(1)"∃/al3,l4.l2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al3a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:l4)) [@a href="cic:/matita/tutorial/chapter5/memb_exists.def(5)"memb_exists/a @sub //] -* #l3 * #l4 #eql2 >eql2 >a href="cic:/matita/basics/list/length_append.def(2)"length_append/a normalize -applyS a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a <a href="cic:/matita/basics/list/length_append.def(2)"length_append/a @Hind [@(a href="cic:/matita/basics/bool/andb_true_r.def(4)"andb_true_r/a … unique)] +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 (a href="cic:/matita/tutorial/chapter5/memb_append.def(5)"memb_append/a … (sub x (a href="cic:/matita/basics/bool/orb_true_r2.def(3)"orb_true_r2/a … membx))) - [#membxl3 @a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"memb_append_l1/a // - |#membxal4 cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … membxal4) - [#eqxa @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a lapply (a href="cic:/matita/basics/bool/andb_true_l.def(4)"andb_true_l/a … unique) - <(\P eqxa) >membx normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ |#membxl4 @a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"memb_append_l2/a // +cases (memb_append … (sub x (orb_true_r2 … membx))) + [#membxl3 @memb_append_l1 // + |#membxal4 cases (orb_true_l … membxal4) + [#eqxa @False_ind lapply (andb_true_l … unique) + <(\P eqxa) >membx normalize /span class="autotactic"2span class="autotrace" trace absurd/span/span/ |#membxl4 @memb_append_l2 // ] ] qed. lemma sublist_unique_append_l1: - ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2). -#S #l1 elim l1 normalize [#l2 #S #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] + ∀S,l1,l2. sublist S l1 (unique_append S l1 l2). +#S #l1 elim l1 normalize [#l2 #S #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/] #x #tl #Hind #l2 #a -normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) #eqax >eqax -[<(\P eqax) cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2))) - [#H >H normalize // | #H >H normalize >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … a)) //] -|cases (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2)) normalize +normalize cases (true_or_false … (a==x)) #eqax >eqax +[<(\P eqax) cases (true_or_false (memb S a (unique_append S tl l2))) + [#H >H normalize // | #H >H normalize >(\b (refl … a)) //] +|cases (memb S x (unique_append S tl l2)) normalize [/span class="autotactic"2span class="autotrace" trace /span/span/ |>eqax normalize /span class="autotactic"2span class="autotrace" trace /span/span/] ] qed. lemma sublist_unique_append_l2: - ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l2 (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2). + ∀S,l1,l2. sublist S l2 (unique_append S l1 l2). #S #l1 elim l1 [normalize //] #x #tl #Hind normalize -#l2 #a cases (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2)) normalize -[@Hind | cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x) normalize // @Hind] +#l2 #a cases (memb S x (unique_append S tl l2)) normalize +[@Hind | cases (a==x) normalize // @Hind] qed. lemma decidable_sublist:∀S,l1,l2. - (a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2) a title="logical or" href="cic:/fakeuri.def(1)"∨/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2). + (sublist S l1 l2) ∨ ¬(sublist S l1 l2). #S #l1 #l2 elim l1 - [%1 #a normalize in ⊢ (%→?); #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ + [%1 #a normalize in ⊢ (%→?); #abs @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/ |#a #tl * #subtl - [cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2)) #memba - [%1 whd #x #membx cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … membx) + [cases (true_or_false (memb S a l2)) #memba + [%1 whd #x #membx cases (orb_true_l … membx) [#eqax >(\P eqax) // |@subtl] - |%2 @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … (a href="cic:/matita/basics/bool/eqnot_to_noteq.def(4)"eqnot_to_noteq/a … a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a memba)) #H1 @H1 @a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"memb_hd/a + |%2 @(not_to_not … (eqnot_to_noteq … true memba)) #H1 @H1 @memb_hd ] - |%2 @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … subtl) #H1 #x #H2 @H1 @a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"memb_cons/a // + |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons // ] ] qed. @@ -193,38 +193,38 @@ qed. (********************* filtering *****************) lemma filter_true: ∀S,f,a,l. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/basics/list/filter.def(2)"filter/a S f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → f a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #f #a #l elim l [normalize #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] -#b #tl #Hind cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (f b)) #H + memb S a (filter S f l) = true → f a = true. +#S #f #a #l elim l [normalize #H @False_ind /span class="autotactic"2span class="autotrace" trace absurd/span/span/] +#b #tl #Hind cases (true_or_false (f b)) #H normalize >H normalize [2:@Hind] -cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) #eqab +cases (true_or_false (a==b)) #eqab [#_ >(\P eqab) // | >eqab normalize @Hind] qed. lemma memb_filter_memb: ∀S,f,a,l. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/basics/list/filter.def(2)"filter/a S f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. + memb S a (filter S f l) = true → memb S a l = true. #S #f #a #l elim l [normalize //] #b #tl #Hind normalize (cases (f b)) normalize -cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize // @Hind +cases (a==b) normalize // @Hind qed. -lemma memb_filter: ∀S,f,l,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → -a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a (f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a). -/span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a, a href="cic:/matita/tutorial/chapter5/memb_filter_memb.def(5)"memb_filter_memb/a, a href="cic:/matita/tutorial/chapter5/filter_true.def(5)"filter_true/a/span/span/ qed. +lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true → +memb S x l = true ∧ (f x = true). +/span class="autotactic"3span class="autotrace" trace conj, memb_filter_memb, filter_truespan class="error" title="disambiguation error"/span/span/span/ qed. -lemma memb_filter_l: ∀S,f,x,l. (f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → -a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +lemma memb_filter_l: ∀S,f,x,l. (f x = true) → memb S x l = true → +memb S x (filter ? f l) = true. #S #f #x #l #fx elim l normalize // -#b #tl #Hind cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (xa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) #eqxb - [<(\P eqxb) >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … x)) >fx normalize >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … x)) normalize // +#b #tl #Hind cases (true_or_false (x==b)) #eqxb + [<(\P eqxb) >(\b (refl … x)) >fx normalize >(\b (refl … x)) normalize // |>eqxb cases (f b) normalize [>eqxb normalize @Hind| @Hind] ] qed. (********************* exists *****************) -let rec exists (A:Type[0]) (p:A → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool ≝ match l with -[ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a -| cons h t ⇒ a href="cic:/matita/basics/bool/orb.def(1)"orb/a (p h) (exists A p t) +[ nil ⇒ false +| cons h t ⇒ orb (p h) (exists A p t) ]. \ No newline at end of file diff --git a/weblib/tutorial/chapter9.ma b/weblib/tutorial/chapter9.ma index c9a65f57b..147bad1b9 100644 --- a/weblib/tutorial/chapter9.ma +++ b/weblib/tutorial/chapter9.ma @@ -121,7 +121,7 @@ lemma moves_cons: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)" // qed. lemma moves_left : ∀S,a,w,e. - a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S (wa title="append" href="cic:/fakeuri.def(1)"@/aa title="cons" href="cic:/fakeuri.def(1)"[/aa]) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e)). + a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S (wa title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a])) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e)). #S #a #w elim w // #x #tl #Hind #e >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a // qed. @@ -186,65 +186,65 @@ are final, while 8 and 9 are not). h2Move to pit/h2. We conclude this chapter with a few properties of the move opertions in relation -with the pit state. *). +with the pit state. *) -definition pit_pre ≝ λS.λi.〈blank S (|i|), false〉. +definition pit_pre ≝ λS.λi.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S (a title="forget" href="cic:/fakeuri.def(1)"|/ai|), a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. -(* The following function compute if a symbol in S occurs in a given -item i *). +(* The following function compute the list of characters occurring in a given +item i. *) -let rec occur (S: DeqSet) (i: re S) on i ≝ +let rec occur (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a 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) + [ z ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a ] + | e ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a ] + | s y ⇒ ya title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a] + | o e1 e2 ⇒ a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (occur S e1) (occur S e2) + | c e1 e2 ⇒ a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (occur S e1) (occur S e2) | k e ⇒ occur S e]. (* If a symbol a does not occur in i, then move(i,a) gets to the -pit state. *). +pit state. *) -lemma not_occur_to_pit: ∀S,a.∀i:pitem S. memb S a (occur S (|i|)) ≠ true → - move S a i = pit_pre S i. +lemma not_occur_to_pit: ∀S,a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/ai|)) a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → + a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a i a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. #S #a #i elim i // - [#x normalize cases (a==x) normalize // #H @False_ind /2/ - |#i1 #i2 #Hind1 #Hind2 #H >move_cat - >Hind1 [2:@(not_to_not … H) #H1 @sublist_unique_append_l1 //] - >Hind2 [2:@(not_to_not … H) #H1 @sublist_unique_append_l2 //] // - |#i1 #i2 #Hind1 #Hind2 #H >move_plus - >Hind1 [2:@(not_to_not … H) #H1 @sublist_unique_append_l1 //] - >Hind2 [2:@(not_to_not … H) #H1 @sublist_unique_append_l2 //] // - |#i #Hind #H >move_star >Hind // + [#x normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x) normalize // #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ + |#i1 #i2 #Hind1 #Hind2 #H >a href="cic:/matita/tutorial/chapter9/move_cat.def(7)"move_cat/a + >Hind1 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a //] + >Hind2 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"sublist_unique_append_l2/a //] // + |#i1 #i2 #Hind1 #Hind2 #H >a href="cic:/matita/tutorial/chapter9/move_plus.def(7)"move_plus/a + >Hind1 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a //] + >Hind2 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"sublist_unique_append_l2/a //] // + |#i #Hind #H >a href="cic:/matita/tutorial/chapter9/move_star.def(7)"move_star/a >Hind // ] qed. -(* We cannot escape form the pit state. *). +(* We cannot escape form the pit state. *) -lemma move_pit: ∀S,a,i. move S a (\fst (pit_pre S i)) = pit_pre S i. +lemma move_pit: ∀S,a,i. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a 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 // + [#i1 #i2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter9/move_cat.def(7)"move_cat/a >Hind1 >Hind2 // + |#i1 #i2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter9/move_plus.def(7)"move_plus/a >Hind1 >Hind2 // + |#i #Hind >a href="cic:/matita/tutorial/chapter9/move_star.def(7)"move_star/a >Hind // ] qed. -lemma moves_pit: ∀S,w,i. moves S w (pit_pre S i) = pit_pre S i. +lemma moves_pit: ∀S,w,i. a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. #S #w #i elim w // qed. (* If any character in w does not occur in i, then moves(i,w) gets -to the pit state. *). +to the pit state. *) -lemma to_pit: ∀S,w,e. ¬ sublist S w (occur S (|\fst e|)) → - moves S w e = pit_pre S (\fst e). +lemma to_pit: ∀S,w,e. a title="logical not" href="cic:/fakeuri.def(1)"¬/a a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|)) → + a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). #S #w elim w - [#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) + [#e * #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a @H normalize #a #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ + |#a #tl #Hind #e #H cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|)))) + [#Htrue >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a whd in ⊢ (???%); <(a href="cic:/matita/tutorial/chapter9/same_kernel.def(8)"same_kernel/a … a) + @Hind >a href="cic:/matita/tutorial/chapter9/same_kernel.def(8)"same_kernel/a @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 #b #memb cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … memb) [#H2 >(\P H2) // |#H2 @H1 //] - |#Hfalse >moves_cons >not_occur_to_pit // >Hfalse /2/ + |#Hfalse >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a >a href="cic:/matita/tutorial/chapter9/not_occur_to_pit.def(8)"not_occur_to_pit/a // >Hfalse /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/eqnot_to_noteq.def(4)"eqnot_to_noteq/a/span/span/ ] ] qed. \ No newline at end of file