# HG changeset patch # User wenzelm # Date 1466369502 -7200 # Node ID dff40165618cdb405c952a53a51ffac8a0caa17e # Parent 67c38b9ea2fba7a806a8a701a735ef85fa879603 misc tuning and modernization; diff -r 67c38b9ea2fb -r dff40165618c src/HOL/Set.thy --- a/src/HOL/Set.thy Sun Jun 19 17:40:51 2016 +0200 +++ b/src/HOL/Set.thy Sun Jun 19 22:51:42 2016 +0200 @@ -35,7 +35,7 @@ text \Set comprehensions\ syntax - "_Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") + "_Coll" :: "pttrn \ bool \ 'a set" ("(1{_./ _})") translations "{x. P}" \ "CONST Collect (\x. P)" @@ -52,15 +52,15 @@ lemma CollectD: "a \ {x. P x} \ P a" by simp -lemma Collect_cong: "(\x. P x = Q x) ==> {x. P x} = {x. Q x}" +lemma Collect_cong: "(\x. P x = Q x) \ {x. P x} = {x. Q x}" by simp text \ -Simproc for pulling \x=t\ in \{x. \ & x=t & \}\ -to the front (and similarly for \t=x\): + Simproc for pulling \x = t\ in \{x. \ \ x = t \ \}\ + to the front (and similarly for \t = x\): \ -simproc_setup defined_Collect ("{x. P x & Q x}") = \ +simproc_setup defined_Collect ("{x. P x \ Q x}") = \ fn _ => Quantifier1.rearrange_Collect (fn ctxt => resolve_tac ctxt @{thms Collect_cong} 1 THEN @@ -80,8 +80,7 @@ then show ?thesis by simp qed -lemma set_eq_iff: - "A = B \ (\x. x \ A \ x \ B)" +lemma set_eq_iff: "A = B \ (\x. x \ A \ x \ B)" by (auto intro:set_eqI) text \Lifting of predicate class instances\ @@ -89,52 +88,52 @@ instantiation set :: (type) boolean_algebra begin -definition less_eq_set where - "A \ B \ (\x. member x A) \ (\x. member x B)" - -definition less_set where - "A < B \ (\x. member x A) < (\x. member x B)" - -definition inf_set where - "A \ B = Collect ((\x. member x A) \ (\x. member x B))" - -definition sup_set where - "A \ B = Collect ((\x. member x A) \ (\x. member x B))" - -definition bot_set where - "\ = Collect \" - -definition top_set where - "\ = Collect \" - -definition uminus_set where - "- A = Collect (- (\x. member x A))" - -definition minus_set where - "A - B = Collect ((\x. member x A) - (\x. member x B))" - -instance proof -qed (simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def - bot_set_def top_set_def uminus_set_def minus_set_def - less_le_not_le inf_compl_bot sup_compl_top sup_inf_distrib1 diff_eq - set_eqI fun_eq_iff - del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply) +definition less_eq_set + where "A \ B \ (\x. member x A) \ (\x. member x B)" + +definition less_set + where "A < B \ (\x. member x A) < (\x. member x B)" + +definition inf_set + where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" + +definition sup_set + where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" + +definition bot_set + where "\ = Collect \" + +definition top_set + where "\ = Collect \" + +definition uminus_set + where "- A = Collect (- (\x. member x A))" + +definition minus_set + where "A - B = Collect ((\x. member x A) - (\x. member x B))" + +instance + by standard + (simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def + bot_set_def top_set_def uminus_set_def minus_set_def + less_le_not_le sup_inf_distrib1 diff_eq set_eqI fun_eq_iff + del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply) end text \Set enumerations\ -abbreviation empty :: "'a set" ("{}") where - "{} \ bot" - -definition insert :: "'a \ 'a set \ 'a set" where - insert_compr: "insert a B = {x. x = a \ x \ B}" +abbreviation empty :: "'a set" ("{}") + where "{} \ bot" + +definition insert :: "'a \ 'a set \ 'a set" + where insert_compr: "insert a B = {x. x = a \ x \ B}" syntax - "_Finset" :: "args => 'a set" ("{(_)}") + "_Finset" :: "args \ 'a set" ("{(_)}") translations - "{x, xs}" == "CONST insert x {xs}" - "{x}" == "CONST insert x {}" + "{x, xs}" \ "CONST insert x {xs}" + "{x}" \ "CONST insert x {}" subsection \Subsets and bounded quantifiers\ @@ -171,28 +170,28 @@ subset_eq ("op <=") and subset_eq ("(_/ <= _)" [51, 51] 50) -definition Ball :: "'a set \ ('a \ bool) \ bool" where - "Ball A P \ (\x. x \ A \ P x)" \ "bounded universal quantifiers" - -definition Bex :: "'a set \ ('a \ bool) \ bool" where - "Bex A P \ (\x. x \ A \ P x)" \ "bounded existential quantifiers" +definition Ball :: "'a set \ ('a \ bool) \ bool" + where "Ball A P \ (\x. x \ A \ P x)" \ "bounded universal quantifiers" + +definition Bex :: "'a set \ ('a \ bool) \ bool" + where "Bex A P \ (\x. x \ A \ P x)" \ "bounded existential quantifiers" syntax (ASCII) - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) - "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST _:_./ _)" [0, 0, 10] 10) + "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3EX _:_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) + "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST _:_./ _)" [0, 0, 10] 10) syntax (input) - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3! _:_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) + "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3! _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3? _:_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3?! _:_./ _)" [0, 0, 10] 10) syntax - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) - "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) + "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) + "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) translations "\x\A. P" \ "CONST Ball A (\x. P)" @@ -201,25 +200,25 @@ "LEAST x:A. P" \ "LEAST x. x \ A \ P" syntax (ASCII output) - "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) - "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) - "_setleAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) - "_setleEx" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) - "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) + "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) + "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<_./ _)" [0, 0, 10] 10) + "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) + "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) + "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) syntax - "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setlessEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleAll" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleEx" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_setleEx1" :: "[idt, 'a, bool] => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) + "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) translations - "\A\B. P" \ "\A. A \ B \ P" - "\A\B. P" \ "\A. A \ B \ P" - "\A\B. P" \ "\A. A \ B \ P" - "\A\B. P" \ "\A. A \ B \ P" - "\!A\B. P" \ "\!A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\A\B. P" \ "\A. A \ B \ P" + "\!A\B. P" \ "\!A. A \ B \ P" print_translation \ let @@ -256,12 +255,13 @@ text \ - \medskip Translate between \{e | x1...xn. P}\ and \{u. EX x1..xn. u = e & P}\; \{y. EX x1..xn. y = e & P}\ is - only translated if \[0..n] subset bvs(e)\. + \<^medskip> + Translate between \{e | x1\xn. P}\ and \{u. \x1\xn. u = e \ P}\; + \{y. \x1\xn. y = e \ P}\ is only translated if \[0..n] \ bvs e\. \ syntax - "_Setcompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") + "_Setcompr" :: "'a \ idts \ bool \ 'a set" ("(1{_ |/_./ _})") parse_translation \ let @@ -319,31 +319,29 @@ in [(@{const_syntax Collect}, setcompr_tr')] end; \ -simproc_setup defined_Bex ("EX x:A. P x & Q x") = \ +simproc_setup defined_Bex ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_bex (fn ctxt => unfold_tac ctxt @{thms Bex_def} THEN Quantifier1.prove_one_point_ex_tac ctxt) \ -simproc_setup defined_All ("ALL x:A. P x --> Q x") = \ +simproc_setup defined_All ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_ball (fn ctxt => unfold_tac ctxt @{thms Ball_def} THEN Quantifier1.prove_one_point_all_tac ctxt) \ -lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x" +lemma ballI [intro!]: "(\x. x \ A \ P x) \ \x\A. P x" by (simp add: Ball_def) lemmas strip = impI allI ballI -lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x" +lemma bspec [dest?]: "\x\A. P x \ x \ A \ P x" by (simp add: Ball_def) -text \ - Gives better instantiation for bound: -\ +text \Gives better instantiation for bound:\ setup \ map_theory_claset (fn ctxt => @@ -353,98 +351,91 @@ ML \ structure Simpdata = struct - -open Simpdata; - -val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; - + open Simpdata; + val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; end; open Simpdata; \ -declaration \fn _ => - Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs)) -\ - -lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" - by (unfold Ball_def) blast - -lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x" - \ \Normally the best argument order: @{prop "P x"} constrains the - choice of @{prop "x:A"}.\ - by (unfold Bex_def) blast - -lemma rev_bexI [intro?]: "x:A ==> P x ==> EX x:A. P x" - \ \The best argument order when there is only one @{prop "x:A"}.\ - by (unfold Bex_def) blast - -lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x" - by (unfold Bex_def) blast - -lemma bexE [elim!]: "EX x:A. P x ==> (!!x. x:A ==> P x ==> Q) ==> Q" - by (unfold Bex_def) blast - -lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)" +declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))\ + +lemma ballE [elim]: "\x\A. P x \ (P x \ Q) \ (x \ A \ Q) \ Q" + unfolding Ball_def by blast + +lemma bexI [intro]: "P x \ x \ A \ \x\A. P x" + \ \Normally the best argument order: \P x\ constrains the choice of \x \ A\.\ + unfolding Bex_def by blast + +lemma rev_bexI [intro?]: "x \ A \ P x \ \x\A. P x" + \ \The best argument order when there is only one \x \ A\.\ + unfolding Bex_def by blast + +lemma bexCI: "(\x\A. \ P x \ P a) \ a \ A \ \x\A. P x" + unfolding Bex_def by blast + +lemma bexE [elim!]: "\x\A. P x \ (\x. x \ A \ P x \ Q) \ Q" + unfolding Bex_def by blast + +lemma ball_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \Trival rewrite rule.\ by (simp add: Ball_def) -lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)" +lemma bex_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \Dual form for existentials.\ by (simp add: Bex_def) -lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)" +lemma bex_triv_one_point1 [simp]: "(\x\A. x = a) \ a \ A" by blast -lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)" +lemma bex_triv_one_point2 [simp]: "(\x\A. a = x) \ a \ A" by blast -lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)" +lemma bex_one_point1 [simp]: "(\x\A. x = a \ P x) \ a \ A \ P a" by blast -lemma bex_one_point2 [simp]: "(EX x:A. a = x & P x) = (a:A & P a)" +lemma bex_one_point2 [simp]: "(\x\A. a = x \ P x) \ a \ A \ P a" by blast -lemma ball_one_point1 [simp]: "(ALL x:A. x = a --> P x) = (a:A --> P a)" +lemma ball_one_point1 [simp]: "(\x\A. x = a \ P x) \ (a \ A \ P a)" by blast -lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" +lemma ball_one_point2 [simp]: "(\x\A. a = x \ P x) \ (a \ A \ P a)" by blast -lemma ball_conj_distrib: - "(\x\A. P x \ Q x) \ ((\x\A. P x) \ (\x\A. Q x))" +lemma ball_conj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast -lemma bex_disj_distrib: - "(\x\A. P x \ Q x) \ ((\x\A. P x) \ (\x\A. Q x))" +lemma bex_disj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast text \Congruence rules\ lemma ball_cong: - "A = B ==> (!!x. x:B ==> P x = Q x) ==> - (ALL x:A. P x) = (ALL x:B. Q x)" + "A = B \ (\x. x \ B \ P x \ Q x) \ + (\x\A. P x) \ (\x\B. Q x)" by (simp add: Ball_def) lemma strong_ball_cong [cong]: - "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> - (ALL x:A. P x) = (ALL x:B. Q x)" + "A = B \ (\x. x \ B =simp=> P x \ Q x) \ + (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Ball_def) lemma bex_cong: - "A = B ==> (!!x. x:B ==> P x = Q x) ==> - (EX x:A. P x) = (EX x:B. Q x)" + "A = B \ (\x. x \ B \ P x \ Q x) \ + (\x\A. P x) \ (\x\B. Q x)" by (simp add: Bex_def cong: conj_cong) lemma strong_bex_cong [cong]: - "A = B ==> (!!x. x:B =simp=> P x = Q x) ==> - (EX x:A. P x) = (EX x:B. Q x)" + "A = B \ (\x. x \ B =simp=> P x \ Q x) \ + (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Bex_def cong: conj_cong) lemma bex1_def: "(\!x\X. P x) \ (\x\X. P x) \ (\x\X. \y\X. P x \ P y \ x = y)" by auto + subsection \Basic operations\ subsubsection \Subsets\ @@ -453,50 +444,45 @@ by (simp add: less_eq_set_def le_fun_def) text \ - \medskip Map the type \'a set => anything\ to just @{typ - 'a}; for overloading constants whose first argument has type @{typ - "'a set"}. + \<^medskip> + Map the type \'a set \ anything\ to just \'a\; for overloading constants + whose first argument has type \'a set\. \ -lemma subsetD [elim, intro?]: "A \ B ==> c \ A ==> c \ B" +lemma subsetD [elim, intro?]: "A \ B \ c \ A \ c \ B" by (simp add: less_eq_set_def le_fun_def) \ \Rule in Modus Ponens style.\ -lemma rev_subsetD [intro?]: "c \ A ==> A \ B ==> c \ B" - \ \The same, with reversed premises for use with \erule\ -- - cf \rev_mp\.\ +lemma rev_subsetD [intro?]: "c \ A \ A \ B \ c \ B" + \ \The same, with reversed premises for use with \erule\ -- cf. \rev_mp\.\ by (rule subsetD) -text \ - \medskip Converts @{prop "A \ B"} to @{prop "x \ A ==> x \ B"}. -\ - -lemma subsetCE [elim]: "A \ B ==> (c \ A ==> P) ==> (c \ B ==> P) ==> P" +lemma subsetCE [elim]: "A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" \ \Classical elimination rule.\ by (auto simp add: less_eq_set_def le_fun_def) -lemma subset_eq: "A \ B = (\x\A. x \ B)" by blast - -lemma contra_subsetD: "A \ B ==> c \ B ==> c \ A" +lemma subset_eq: "A \ B = (\x\A. x \ B)" + by blast + +lemma contra_subsetD: "A \ B \ c \ B \ c \ A" by blast lemma subset_refl: "A \ A" by (fact order_refl) (* already [iff] *) -lemma subset_trans: "A \ B ==> B \ C ==> A \ C" +lemma subset_trans: "A \ B \ B \ C \ A \ C" by (fact order_trans) -lemma set_rev_mp: "x:A ==> A \ B ==> x:B" +lemma set_rev_mp: "x \ A \ A \ B \ x \ B" by (rule subsetD) -lemma set_mp: "A \ B ==> x:A ==> x:B" +lemma set_mp: "A \ B \ x \ A \ x \ B" by (rule subsetD) -lemma subset_not_subset_eq [code]: - "A \ B \ A \ B \ \ B \ A" +lemma subset_not_subset_eq [code]: "A \ B \ A \ B \ \ B \ A" by (fact less_le_not_le) -lemma eq_mem_trans: "a=b ==> b \ A ==> a \ A" +lemma eq_mem_trans: "a = b \ b \ A \ a \ A" by simp lemmas basic_trans_rules [trans] = @@ -505,123 +491,120 @@ subsubsection \Equality\ -lemma subset_antisym [intro!]: "A \ B ==> B \ A ==> A = B" +lemma subset_antisym [intro!]: "A \ B \ B \ A \ A = B" \ \Anti-symmetry of the subset relation.\ by (iprover intro: set_eqI subsetD) -text \ - \medskip Equality rules from ZF set theory -- are they appropriate - here? -\ - -lemma equalityD1: "A = B ==> A \ B" +text \\<^medskip> Equality rules from ZF set theory -- are they appropriate here?\ + +lemma equalityD1: "A = B \ A \ B" by simp -lemma equalityD2: "A = B ==> B \ A" +lemma equalityD2: "A = B \ B \ A" by simp text \ - \medskip Be careful when adding this to the claset as \subset_empty\ is in the simpset: @{prop "A = {}"} goes to @{prop "{} - \ A"} and @{prop "A \ {}"} and then back to @{prop "A = {}"}! + \<^medskip> + Be careful when adding this to the claset as \subset_empty\ is in the + simpset: @{prop "A = {}"} goes to @{prop "{} \ A"} and @{prop "A \ {}"} + and then back to @{prop "A = {}"}! \ -lemma equalityE: "A = B ==> (A \ B ==> B \ A ==> P) ==> P" +lemma equalityE: "A = B \ (A \ B \ B \ A \ P) \ P" by simp -lemma equalityCE [elim]: - "A = B ==> (c \ A ==> c \ B ==> P) ==> (c \ A ==> c \ B ==> P) ==> P" +lemma equalityCE [elim]: "A = B \ (c \ A \ c \ B \ P) \ (c \ A \ c \ B \ P) \ P" by blast -lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)" +lemma eqset_imp_iff: "A = B \ x \ A \ x \ B" by simp -lemma eqelem_imp_iff: "x = y ==> (x : A) = (y : A)" +lemma eqelem_imp_iff: "x = y \ x \ A \ y \ A" by simp subsubsection \The empty set\ -lemma empty_def: - "{} = {x. False}" +lemma empty_def: "{} = {x. False}" by (simp add: bot_set_def bot_fun_def) -lemma empty_iff [simp]: "(c : {}) = False" +lemma empty_iff [simp]: "c \ {} \ False" by (simp add: empty_def) -lemma emptyE [elim!]: "a : {} ==> P" +lemma emptyE [elim!]: "a \ {} \ P" by simp lemma empty_subsetI [iff]: "{} \ A" - \ \One effect is to delete the ASSUMPTION @{prop "{} <= A"}\ + \ \One effect is to delete the ASSUMPTION @{prop "{} \ A"}\ by blast -lemma equals0I: "(!!y. y \ A ==> False) ==> A = {}" +lemma equals0I: "(\y. y \ A \ False) \ A = {}" by blast -lemma equals0D: "A = {} ==> a \ A" - \ \Use for reasoning about disjointness: \A Int B = {}\\ +lemma equals0D: "A = {} \ a \ A" + \ \Use for reasoning about disjointness: \A \ B = {}\\ by blast -lemma ball_empty [simp]: "Ball {} P = True" +lemma ball_empty [simp]: "Ball {} P \ True" by (simp add: Ball_def) -lemma bex_empty [simp]: "Bex {} P = False" +lemma bex_empty [simp]: "Bex {} P \ False" by (simp add: Bex_def) subsubsection \The universal set -- UNIV\ -abbreviation UNIV :: "'a set" where - "UNIV \ top" - -lemma UNIV_def: - "UNIV = {x. True}" +abbreviation UNIV :: "'a set" + where "UNIV \ top" + +lemma UNIV_def: "UNIV = {x. True}" by (simp add: top_set_def top_fun_def) -lemma UNIV_I [simp]: "x : UNIV" +lemma UNIV_I [simp]: "x \ UNIV" by (simp add: UNIV_def) declare UNIV_I [intro] \ \unsafe makes it less likely to cause problems\ -lemma UNIV_witness [intro?]: "EX x. x : UNIV" +lemma UNIV_witness [intro?]: "\x. x \ UNIV" by simp lemma subset_UNIV: "A \ UNIV" by (fact top_greatest) (* already simp *) text \ - \medskip Eta-contracting these two rules (to remove \P\) - causes them to be ignored because of their interaction with - congruence rules. + \<^medskip> + Eta-contracting these two rules (to remove \P\) causes them + to be ignored because of their interaction with congruence rules. \ -lemma ball_UNIV [simp]: "Ball UNIV P = All P" +lemma ball_UNIV [simp]: "Ball UNIV P \ All P" by (simp add: Ball_def) -lemma bex_UNIV [simp]: "Bex UNIV P = Ex P" +lemma bex_UNIV [simp]: "Bex UNIV P \ Ex P" by (simp add: Bex_def) lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" by auto -lemma UNIV_not_empty [iff]: "UNIV ~= {}" +lemma UNIV_not_empty [iff]: "UNIV \ {}" by (blast elim: equalityE) lemma empty_not_UNIV[simp]: "{} \ UNIV" -by blast + by blast + subsubsection \The Powerset operator -- Pow\ -definition Pow :: "'a set => 'a set set" where - Pow_def: "Pow A = {B. B \ A}" - -lemma Pow_iff [iff]: "(A \ Pow B) = (A \ B)" +definition Pow :: "'a set \ 'a set set" + where Pow_def: "Pow A = {B. B \ A}" + +lemma Pow_iff [iff]: "A \ Pow B \ A \ B" by (simp add: Pow_def) -lemma PowI: "A \ B ==> A \ Pow B" +lemma PowI: "A \ B \ A \ Pow B" by (simp add: Pow_def) -lemma PowD: "A \ Pow B ==> A \ B" +lemma PowD: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma Pow_bottom: "{} \ Pow B" @@ -636,23 +619,25 @@ subsubsection \Set complement\ -lemma Compl_iff [simp]: "(c \ -A) = (c \ A)" +lemma Compl_iff [simp]: "c \ - A \ c \ A" by (simp add: fun_Compl_def uminus_set_def) -lemma ComplI [intro!]: "(c \ A ==> False) ==> c \ -A" +lemma ComplI [intro!]: "(c \ A \ False) \ c \ - A" by (simp add: fun_Compl_def uminus_set_def) blast text \ - \medskip This form, with negated conclusion, works well with the - Classical prover. Negated assumptions behave like formulae on the - right side of the notional turnstile ...\ - -lemma ComplD [dest!]: "c : -A ==> c~:A" + \<^medskip> + This form, with negated conclusion, works well with the Classical prover. + Negated assumptions behave like formulae on the right side of the + notional turnstile \dots +\ + +lemma ComplD [dest!]: "c \ - A \ c \ A" by simp lemmas ComplE = ComplD [elim_format] -lemma Compl_eq: "- A = {x. ~ x : A}" +lemma Compl_eq: "- A = {x. \ x \ A}" by blast @@ -664,23 +649,22 @@ notation (ASCII) inter (infixl "Int" 70) -lemma Int_def: - "A \ B = {x. x \ A \ x \ B}" +lemma Int_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: inf_set_def inf_fun_def) -lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" - by (unfold Int_def) blast - -lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" +lemma Int_iff [simp]: "c \ A \ B \ c \ A \ c \ B" + unfolding Int_def by blast + +lemma IntI [intro!]: "c \ A \ c \ B \ c \ A \ B" by simp -lemma IntD1: "c : A Int B ==> c:A" +lemma IntD1: "c \ A \ B \ c \ A" by simp -lemma IntD2: "c : A Int B ==> c:B" +lemma IntD2: "c \ A \ B \ c \ B" by simp -lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" +lemma IntE [elim!]: "c \ A \ B \ (c \ A \ c \ B \ P) \ P" by simp lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" @@ -695,29 +679,25 @@ notation (ASCII) union (infixl "Un" 65) -lemma Un_def: - "A \ B = {x. x \ A \ x \ B}" +lemma Un_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: sup_set_def sup_fun_def) -lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" - by (unfold Un_def) blast - -lemma UnI1 [elim?]: "c:A ==> c : A Un B" - by simp - -lemma UnI2 [elim?]: "c:B ==> c : A Un B" +lemma Un_iff [simp]: "c \ A \ B \ c \ A \ c \ B" + unfolding Un_def by blast + +lemma UnI1 [elim?]: "c \ A \ c \ A \ B" by simp -text \ - \medskip Classical introduction rule: no commitment to @{prop A} vs - @{prop B}. -\ - -lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B" +lemma UnI2 [elim?]: "c \ B \ c \ A \ B" + by simp + +text \\<^medskip> Classical introduction rule: no commitment to @{prop A} vs @{prop B}.\ + +lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by auto -lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P" - by (unfold Un_def) blast +lemma UnE [elim!]: "c \ A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" + unfolding Un_def by blast lemma insert_def: "insert a B = {x. x = a} \ B" by (simp add: insert_compr Un_def) @@ -728,109 +708,110 @@ subsubsection \Set difference\ -lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)" +lemma Diff_iff [simp]: "c \ A - B \ c \ A \ c \ B" by (simp add: minus_set_def fun_diff_def) -lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B" +lemma DiffI [intro!]: "c \ A \ c \ B \ c \ A - B" by simp -lemma DiffD1: "c : A - B ==> c : A" +lemma DiffD1: "c \ A - B \ c \ A" by simp -lemma DiffD2: "c : A - B ==> c : B ==> P" +lemma DiffD2: "c \ A - B \ c \ B \ P" by simp -lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P" +lemma DiffE [elim!]: "c \ A - B \ (c \ A \ c \ B \ P) \ P" by simp -lemma set_diff_eq: "A - B = {x. x : A & ~ x : B}" by blast - -lemma Compl_eq_Diff_UNIV: "-A = (UNIV - A)" -by blast +lemma set_diff_eq: "A - B = {x. x \ A \ x \ B}" + by blast + +lemma Compl_eq_Diff_UNIV: "- A = (UNIV - A)" + by blast subsubsection \Augmenting a set -- @{const insert}\ -lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)" - by (unfold insert_def) blast - -lemma insertI1: "a : insert a B" +lemma insert_iff [simp]: "a \ insert b A \ a = b \ a \ A" + unfolding insert_def by blast + +lemma insertI1: "a \ insert a B" by simp -lemma insertI2: "a : B ==> a : insert b B" +lemma insertI2: "a \ B \ a \ insert b B" by simp -lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P" - by (unfold insert_def) blast - -lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B" +lemma insertE [elim!]: "a \ insert b A \ (a = b \ P) \ (a \ A \ P) \ P" + unfolding insert_def by blast + +lemma insertCI [intro!]: "(a \ B \ a = b) \ a \ insert b B" \ \Classical introduction rule.\ by auto -lemma subset_insert_iff: "(A \ insert x B) = (if x:A then A - {x} \ B else A \ B)" +lemma subset_insert_iff: "A \ insert x B \ (if x \ A then A - {x} \ B else A \ B)" by auto lemma set_insert: assumes "x \ A" obtains B where "A = insert x B" and "x \ B" proof - from assms show "A = insert x (A - {x})" by blast -next + show "A = insert x (A - {x})" using assms by blast show "x \ A - {x}" by blast qed -lemma insert_ident: "x ~: A ==> x ~: B ==> (insert x A = insert x B) = (A = B)" -by auto - -lemma insert_eq_iff: assumes "a \ A" "b \ B" -shows "insert a A = insert b B \ - (if a=b then A=B else \C. A = insert b C \ b \ C \ B = insert a C \ a \ C)" - (is "?L \ ?R") +lemma insert_ident: "x \ A \ x \ B \ insert x A = insert x B \ A = B" + by auto + +lemma insert_eq_iff: + assumes "a \ A" "b \ B" + shows "insert a A = insert b B \ + (if a = b then A = B else \C. A = insert b C \ b \ C \ B = insert a C \ a \ C)" + (is "?L \ ?R") proof - assume ?L - show ?R - proof cases - assume "a=b" with assms \?L\ show ?R by (simp add: insert_ident) + show ?R if ?L + proof (cases "a = b") + case True + with assms \?L\ show ?R + by (simp add: insert_ident) next - assume "a\b" + case False let ?C = "A - {b}" have "A = insert b ?C \ b \ ?C \ B = insert a ?C \ a \ ?C" - using assms \?L\ \a\b\ by auto - thus ?R using \a\b\ by auto + using assms \?L\ \a \ b\ by auto + then show ?R using \a \ b\ by auto qed -next - assume ?R thus ?L by (auto split: if_splits) + show ?L if ?R + using that by (auto split: if_splits) qed lemma insert_UNIV: "insert x UNIV = UNIV" -by auto + by auto + subsubsection \Singletons, using insert\ -lemma singletonI [intro!]: "a : {a}" - \ \Redundant? But unlike \insertCI\, it proves the subgoal immediately!\ +lemma singletonI [intro!]: "a \ {a}" + \ \Redundant? But unlike \insertCI\, it proves the subgoal immediately!\ by (rule insertI1) -lemma singletonD [dest!]: "b : {a} ==> b = a" +lemma singletonD [dest!]: "b \ {a} \ b = a" by blast lemmas singletonE = singletonD [elim_format] -lemma singleton_iff: "(b : {a}) = (b = a)" +lemma singleton_iff: "b \ {a} \ b = a" by blast -lemma singleton_inject [dest!]: "{a} = {b} ==> a = b" +lemma singleton_inject [dest!]: "{a} = {b} \ a = b" by blast -lemma singleton_insert_inj_eq [iff]: - "({b} = insert a A) = (a = b & A \ {b})" +lemma singleton_insert_inj_eq [iff]: "{b} = insert a A \ a = b \ A \ {b}" by blast -lemma singleton_insert_inj_eq' [iff]: - "(insert a A = {b}) = (a = b & A \ {b})" +lemma singleton_insert_inj_eq' [iff]: "insert a A = {b} \ a = b \ A \ {b}" by blast -lemma subset_singletonD: "A \ {x} ==> A = {} | A = {x}" +lemma subset_singletonD: "A \ {x} \ A = {} \ A = {x}" by fast lemma subset_singleton_iff: "X \ {a} \ X = {} \ X = {a}" @@ -842,73 +823,59 @@ lemma singleton_conv2 [simp]: "{x. a = x} = {a}" by blast -lemma Diff_single_insert: "A - {x} \ B ==> A \ insert x B" +lemma Diff_single_insert: "A - {x} \ B \ A \ insert x B" by blast -lemma subset_Diff_insert: "A \ B - (insert x C) \ A \ B - C \ x \ A" +lemma subset_Diff_insert: "A \ B - insert x C \ A \ B - C \ x \ A" by blast -lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)" +lemma doubleton_eq_iff: "{a, b} = {c, d} \ a = c \ b = d \ a = d & b = c" by (blast elim: equalityE) -lemma Un_singleton_iff: - "(A \ B = {x}) = (A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x})" -by auto - -lemma singleton_Un_iff: - "({x} = A \ B) = (A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x})" -by auto +lemma Un_singleton_iff: "A \ B = {x} \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" + by auto + +lemma singleton_Un_iff: "{x} = A \ B \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" + by auto subsubsection \Image of a set under a function\ -text \ - Frequently @{term b} does not have the syntactic form of @{term "f x"}. -\ - -definition image :: "('a \ 'b) \ 'a set \ 'b set" (infixr "`" 90) -where - "f ` A = {y. \x\A. y = f x}" - -lemma image_eqI [simp, intro]: - "b = f x \ x \ A \ b \ f ` A" - by (unfold image_def) blast - -lemma imageI: - "x \ A \ f x \ f ` A" +text \Frequently \b\ does not have the syntactic form of \f x\.\ + +definition image :: "('a \ 'b) \ 'a set \ 'b set" (infixr "`" 90) + where "f ` A = {y. \x\A. y = f x}" + +lemma image_eqI [simp, intro]: "b = f x \ x \ A \ b \ f ` A" + unfolding image_def by blast + +lemma imageI: "x \ A \ f x \ f ` A" by (rule image_eqI) (rule refl) -lemma rev_image_eqI: - "x \ A \ b = f x \ b \ f ` A" - \ \This version's more effective when we already have the - required @{term x}.\ +lemma rev_image_eqI: "x \ A \ b = f x \ b \ f ` A" + \ \This version's more effective when we already have the required \x\.\ by (rule image_eqI) lemma imageE [elim!]: - assumes "b \ (\x. f x) ` A" \ \The eta-expansion gives variable-name preservation.\ + assumes "b \ (\x. f x) ` A" \ \The eta-expansion gives variable-name preservation.\ obtains x where "b = f x" and "x \ A" - using assms by (unfold image_def) blast - -lemma Compr_image_eq: - "{x \ f ` A. P x} = f ` {x \ A. P (f x)}" + using assms unfolding image_def by blast + +lemma Compr_image_eq: "{x \ f ` A. P x} = f ` {x \ A. P (f x)}" by auto -lemma image_Un: - "f ` (A \ B) = f ` A \ f ` B" +lemma image_Un: "f ` (A \ B) = f ` A \ f ` B" by blast -lemma image_iff: - "z \ f ` A \ (\x\A. z = f x)" +lemma image_iff: "z \ f ` A \ (\x\A. z = f x)" by blast -lemma image_subsetI: - "(\x. x \ A \ f x \ B) \ f ` A \ B" +lemma image_subsetI: "(\x. x \ A \ f x \ B) \ f ` A \ B" \ \Replaces the three steps \subsetI\, \imageE\, \hypsubst\, but breaks too many existing proofs.\ by blast -lemma image_subset_iff: - "f ` A \ B \ (\x\A. f x \ B)" +lemma image_subset_iff: "f ` A \ B \ (\x\A. f x \ B)" \ \This rewrite rule would confuse users if made default.\ by blast @@ -921,68 +888,53 @@ ultimately show thesis by (blast intro: that) qed -lemma subset_image_iff: - "B \ f ` A \ (\AA\A. B = f ` AA)" +lemma subset_image_iff: "B \ f ` A \ (\AA\A. B = f ` AA)" by (blast elim: subset_imageE) -lemma image_ident [simp]: - "(\x. x) ` Y = Y" +lemma image_ident [simp]: "(\x. x) ` Y = Y" by blast -lemma image_empty [simp]: - "f ` {} = {}" +lemma image_empty [simp]: "f ` {} = {}" by blast -lemma image_insert [simp]: - "f ` insert a B = insert (f a) (f ` B)" +lemma image_insert [simp]: "f ` insert a B = insert (f a) (f ` B)" by blast -lemma image_constant: - "x \ A \ (\x. c) ` A = {c}" +lemma image_constant: "x \ A \ (\x. c) ` A = {c}" by auto -lemma image_constant_conv: - "(\x. c) ` A = (if A = {} then {} else {c})" +lemma image_constant_conv: "(\x. c) ` A = (if A = {} then {} else {c})" by auto -lemma image_image: - "f ` (g ` A) = (\x. f (g x)) ` A" +lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" by blast -lemma insert_image [simp]: - "x \ A ==> insert (f x) (f ` A) = f ` A" +lemma insert_image [simp]: "x \ A \ insert (f x) (f ` A) = f ` A" by blast -lemma image_is_empty [iff]: - "f ` A = {} \ A = {}" +lemma image_is_empty [iff]: "f ` A = {} \ A = {}" by blast -lemma empty_is_image [iff]: - "{} = f ` A \ A = {}" +lemma empty_is_image [iff]: "{} = f ` A \ A = {}" by blast -lemma image_Collect: - "f ` {x. P x} = {f x | x. P x}" - \ \NOT suitable as a default simprule: the RHS isn't simpler than the LHS, +lemma image_Collect: "f ` {x. P x} = {f x | x. P x}" + \ \NOT suitable as a default simp rule: the RHS isn't simpler than the LHS, with its implicit quantifier and conjunction. Also image enjoys better equational properties than does the RHS.\ by blast lemma if_image_distrib [simp]: - "(\x. if P x then f x else g x) ` S - = (f ` (S \ {x. P x})) \ (g ` (S \ {x. \ P x}))" + "(\x. if P x then f x else g x) ` S = f ` (S \ {x. P x}) \ g ` (S \ {x. \ P x})" by auto -lemma image_cong: - "M = N \ (\x. x \ N \ f x = g x) \ f ` M = g ` N" +lemma image_cong: "M = N \ (\x. x \ N \ f x = g x) \ f ` M = g ` N" by (simp add: image_def) -lemma image_Int_subset: - "f ` (A \ B) \ f ` A \ f ` B" +lemma image_Int_subset: "f ` (A \ B) \ f ` A \ f ` B" by blast -lemma image_diff_subset: - "f ` A - f ` B \ f ` (A - B)" +lemma image_diff_subset: "f ` A - f ` B \ f ` (A - B)" by blast lemma Setcompr_eq_image: "{f x | x. x \ A} = f ` A" @@ -991,78 +943,67 @@ lemma setcompr_eq_image: "{f x |x. P x} = f ` {x. P x}" by auto -lemma ball_imageD: - assumes "\x\f ` A. P x" - shows "\x\A. P (f x)" - using assms by simp - -lemma bex_imageD: - assumes "\x\f ` A. P x" - shows "\x\A. P (f x)" - using assms by auto - -lemma image_add_0 [simp]: "op+ (0::'a::comm_monoid_add) ` S = S" +lemma ball_imageD: "\x\f ` A. P x \ \x\A. P (f x)" + by simp + +lemma bex_imageD: "\x\f ` A. P x \ \x\A. P (f x)" + by auto + +lemma image_add_0 [simp]: "op + (0::'a::comm_monoid_add) ` S = S" by auto -text \ - \medskip Range of a function -- just a translation for image! -\ - -abbreviation range :: "('a \ 'b) \ 'b set" -where \ "of function" - "range f \ f ` UNIV" - -lemma range_eqI: - "b = f x \ b \ range f" +text \\<^medskip> Range of a function -- just an abbreviation for image!\ + +abbreviation range :: "('a \ 'b) \ 'b set" \ "of function" + where "range f \ f ` UNIV" + +lemma range_eqI: "b = f x \ b \ range f" by simp -lemma rangeI: - "f x \ range f" +lemma rangeI: "f x \ range f" by simp -lemma rangeE [elim?]: - "b \ range (\x. f x) \ (\x. b = f x \ P) \ P" +lemma rangeE [elim?]: "b \ range (\x. f x) \ (\x. b = f x \ P) \ P" by (rule imageE) -lemma full_SetCompr_eq: - "{u. \x. u = f x} = range f" +lemma full_SetCompr_eq: "{u. \x. u = f x} = range f" by auto -lemma range_composition: - "range (\x. f (g x)) = f ` range g" +lemma range_composition: "range (\x. f (g x)) = f ` range g" by auto subsubsection \Some rules with \if\\ -text\Elimination of \{x. \ & x=t & \}\.\ - -lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})" +text \Elimination of \{x. \ \ x = t \ \}\.\ + +lemma Collect_conv_if: "{x. x = a \ P x} = (if P a then {a} else {})" by auto -lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})" +lemma Collect_conv_if2: "{x. a = x \ P x} = (if P a then {a} else {})" by auto text \ Rewrite rules for boolean case-splitting: faster than \if_split [split]\. \ -lemma if_split_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))" +lemma if_split_eq1: "(if Q then x else y) = b \ (Q \ x = b) \ (\ Q \ y = b)" by (rule if_split) -lemma if_split_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))" +lemma if_split_eq2: "a = (if Q then x else y) \ (Q \ a = x) \ (\ Q \ a = y)" by (rule if_split) text \ - Split ifs on either side of the membership relation. Not for \[simp]\ -- can cause goals to blow up! + Split ifs on either side of the membership relation. + Not for \[simp]\ -- can cause goals to blow up! \ -lemma if_split_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))" +lemma if_split_mem1: "(if Q then x else y) \ b \ (Q \ x \ b) \ (\ Q \ y \ b)" by (rule if_split) -lemma if_split_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))" - by (rule if_split [where P="%S. a : S"]) +lemma if_split_mem2: "(a \ (if Q then x else y)) \ (Q \ a \ x) \ (\ Q \ a \ y)" + by (rule if_split [where P = "\S. a \ S"]) lemmas split_ifs = if_bool_eq_conj if_split_eq1 if_split_eq2 if_split_mem1 if_split_mem2 @@ -1080,58 +1021,48 @@ subsubsection \The ``proper subset'' relation\ -lemma psubsetI [intro!]: "A \ B ==> A \ B ==> A \ B" - by (unfold less_le) blast - -lemma psubsetE [elim!]: - "[|A \ B; [|A \ B; ~ (B\A)|] ==> R|] ==> R" - by (unfold less_le) blast +lemma psubsetI [intro!]: "A \ B \ A \ B \ A \ B" + unfolding less_le by blast + +lemma psubsetE [elim!]: "A \ B \ (A \ B \ \ B \ A \ R) \ R" + unfolding less_le by blast lemma psubset_insert_iff: - "(A \ insert x B) = (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" + "A \ insert x B \ (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" by (auto simp add: less_le subset_insert_iff) -lemma psubset_eq: "(A \ B) = (A \ B & A \ B)" +lemma psubset_eq: "A \ B \ A \ B \ A \ B" by (simp only: less_le) -lemma psubset_imp_subset: "A \ B ==> A \ B" +lemma psubset_imp_subset: "A \ B \ A \ B" by (simp add: psubset_eq) -lemma psubset_trans: "[| A \ B; B \ C |] ==> A \ C" -apply (unfold less_le) -apply (auto dest: subset_antisym) -done - -lemma psubsetD: "[| A \ B; c \ A |] ==> c \ B" -apply (unfold less_le) -apply (auto dest: subsetD) -done - -lemma psubset_subset_trans: "A \ B ==> B \ C ==> A \ C" +lemma psubset_trans: "A \ B \ B \ C \ A \ C" + unfolding less_le by (auto dest: subset_antisym) + +lemma psubsetD: "A \ B \ c \ A \ c \ B" + unfolding less_le by (auto dest: subsetD) + +lemma psubset_subset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) -lemma subset_psubset_trans: "A \ B ==> B \ C ==> A \ C" +lemma subset_psubset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) -lemma psubset_imp_ex_mem: "A \ B ==> \b. b \ (B - A)" - by (unfold less_le) blast - -lemma atomize_ball: - "(!!x. x \ A ==> P x) == Trueprop (\x\A. P x)" +lemma psubset_imp_ex_mem: "A \ B \ \b. b \ B - A" + unfolding less_le by blast + +lemma atomize_ball: "(\x. x \ A \ P x) \ Trueprop (\x\A. P x)" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball -lemma image_Pow_mono: - assumes "f ` A \ B" - shows "image f ` Pow A \ Pow B" - using assms by blast - -lemma image_Pow_surj: - assumes "f ` A = B" - shows "image f ` Pow A = Pow B" - using assms by (blast elim: subset_imageE) +lemma image_Pow_mono: "f ` A \ B \ image f ` Pow A \ Pow B" + by blast + +lemma image_Pow_surj: "f ` A = B \ image f ` Pow A = Pow B" + by (blast elim: subset_imageE) subsubsection \Derived rules involving subsets.\ @@ -1144,11 +1075,11 @@ lemma subset_insertI2: "A \ B \ A \ insert b B" by blast -lemma subset_insert: "x \ A ==> (A \ insert x B) = (A \ B)" +lemma subset_insert: "x \ A \ A \ insert x B \ A \ B" by blast -text \\medskip Finite Union -- the least upper bound of two sets.\ +text \\<^medskip> Finite Union -- the least upper bound of two sets.\ lemma Un_upper1: "A \ A \ B" by (fact sup_ge1) @@ -1156,11 +1087,11 @@ lemma Un_upper2: "B \ A \ B" by (fact sup_ge2) -lemma Un_least: "A \ C ==> B \ C ==> A \ B \ C" +lemma Un_least: "A \ C \ B \ C \ A \ B \ C" by (fact sup_least) -text \\medskip Finite Intersection -- the greatest lower bound of two sets.\ +text \\<^medskip> Finite Intersection -- the greatest lower bound of two sets.\ lemma Int_lower1: "A \ B \ A" by (fact inf_le1) @@ -1168,17 +1099,17 @@ lemma Int_lower2: "A \ B \ B" by (fact inf_le2) -lemma Int_greatest: "C \ A ==> C \ B ==> C \ A \ B" +lemma Int_greatest: "C \ A \ C \ B \ C \ A \ B" by (fact inf_greatest) -text \\medskip Set difference.\ +text \\<^medskip> Set difference.\ lemma Diff_subset: "A - B \ A" by blast -lemma Diff_subset_conv: "(A - B \ C) = (A \ B \ C)" -by blast +lemma Diff_subset_conv: "A - B \ C \ A \ B \ C" + by blast subsubsection \Equalities involving union, intersection, inclusion, etc.\ @@ -1189,49 +1120,47 @@ \ \supersedes \Collect_False_empty\\ by auto -lemma subset_empty [simp]: "(A \ {}) = (A = {})" +lemma subset_empty [simp]: "A \ {} \ A = {}" by (fact bot_unique) lemma not_psubset_empty [iff]: "\ (A < {})" by (fact not_less_bot) (* FIXME: already simp *) -lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\x. \ P x)" -by blast - -lemma empty_Collect_eq [simp]: "({} = Collect P) = (\x. \ P x)" -by blast +lemma Collect_empty_eq [simp]: "Collect P = {} \ (\x. \ P x)" + by blast + +lemma empty_Collect_eq [simp]: "{} = Collect P \ (\x. \ P x)" + by blast lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" by blast -lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \ {x. Q x}" +lemma Collect_disj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast -lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \ {x. Q x}" +lemma Collect_imp_eq: "{x. P x \ Q x} = - {x. P x} \ {x. Q x}" by blast -lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \ {x. Q x}" +lemma Collect_conj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_mono_iff: "Collect P \ Collect Q \ (\x. P x \ Q x)" by blast -text \\medskip \insert\.\ - -lemma insert_is_Un: "insert a A = {a} Un A" - \ \NOT SUITABLE FOR REWRITING since \{a} == insert a {}\\ +text \\<^medskip> \insert\.\ + +lemma insert_is_Un: "insert a A = {a} \ A" + \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a {}\\ by blast lemma insert_not_empty [simp]: "insert a A \ {}" - by blast - -lemmas empty_not_insert = insert_not_empty [symmetric] -declare empty_not_insert [simp] - -lemma insert_absorb: "a \ A ==> insert a A = A" + and empty_not_insert [simp]: "{} \ insert a A" + by blast+ + +lemma insert_absorb: "a \ A \ insert a A = A" \ \\[simp]\ causes recursive calls when there are nested inserts\ - \ \with \emph{quadratic} running time\ + \ \with \<^emph>\quadratic\ running time\ by blast lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" @@ -1240,32 +1169,31 @@ lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" by blast -lemma insert_subset [simp]: "(insert x A \ B) = (x \ B & A \ B)" +lemma insert_subset [simp]: "insert x A \ B \ x \ B \ A \ B" by blast -lemma mk_disjoint_insert: "a \ A ==> \B. A = insert a B & a \ B" +lemma mk_disjoint_insert: "a \ A \ \B. A = insert a B \ a \ B" \ \use new \B\ rather than \A - {a}\ to avoid infinite unfolding\ - apply (rule_tac x = "A - {a}" in exI, blast) - done - -lemma insert_Collect: "insert a (Collect P) = {u. u \ a --> P u}" + by (rule exI [where x = "A - {a}"]) blast + +lemma insert_Collect: "insert a (Collect P) = {u. u \ a \ P u}" by auto -lemma insert_inter_insert[simp]: "insert a A \ insert a B = insert a (A \ B)" +lemma insert_inter_insert [simp]: "insert a A \ insert a B = insert a (A \ B)" by blast lemma insert_disjoint [simp]: - "(insert a A \ B = {}) = (a \ B \ A \ B = {})" - "({} = insert a A \ B) = (a \ B \ {} = A \ B)" + "insert a A \ B = {} \ a \ B \ A \ B = {}" + "{} = insert a A \ B \ a \ B \ {} = A \ B" by auto lemma disjoint_insert [simp]: - "(B \ insert a A = {}) = (a \ B \ B \ A = {})" - "({} = A \ insert b B) = (b \ A \ {} = A \ B)" + "B \ insert a A = {} \ a \ B \ B \ A = {}" + "{} = A \ insert b B \ b \ A \ {} = A \ B" by auto -text \\medskip \Int\\ +text \\<^medskip> \Int\\ lemma Int_absorb: "A \ A = A" by (fact inf_idem) (* already simp *) @@ -1285,10 +1213,10 @@ lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute \ \Intersection is an AC-operator\ -lemma Int_absorb1: "B \ A ==> A \ B = B" +lemma Int_absorb1: "B \ A \ A \ B = B" by (fact inf_absorb2) -lemma Int_absorb2: "A \ B ==> A \ B = A" +lemma Int_absorb2: "A \ B \ A \ B = A" by (fact inf_absorb1) lemma Int_empty_left: "{} \ B = {}" @@ -1297,10 +1225,10 @@ lemma Int_empty_right: "A \ {} = {}" by (fact inf_bot_right) (* already simp *) -lemma disjoint_eq_subset_Compl: "(A \ B = {}) = (A \ -B)" +lemma disjoint_eq_subset_Compl: "A \ B = {} \ A \ - B" by blast -lemma disjoint_iff_not_equal: "(A \ B = {}) = (\x\A. \y\B. x \ y)" +lemma disjoint_iff_not_equal: "A \ B = {} \ (\x\A. \y\B. x \ y)" by blast lemma Int_UNIV_left: "UNIV \ B = B" @@ -1315,17 +1243,17 @@ lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact inf_sup_distrib2) -lemma Int_UNIV [simp]: "(A \ B = UNIV) = (A = UNIV & B = UNIV)" +lemma Int_UNIV [simp]: "A \ B = UNIV \ A = UNIV \ B = UNIV" by (fact inf_eq_top_iff) (* already simp *) -lemma Int_subset_iff [simp]: "(C \ A \ B) = (C \ A & C \ B)" +lemma Int_subset_iff [simp]: "C \ A \ B \ C \ A \ C \ B" by (fact le_inf_iff) -lemma Int_Collect: "(x \ A \ {x. P x}) = (x \ A & P x)" +lemma Int_Collect: "x \ A \ {x. P x} \ x \ A \ P x" by blast -text \\medskip \Un\.\ +text \\<^medskip> \Un\.\ lemma Un_absorb: "A \ A = A" by (fact sup_idem) (* already simp *) @@ -1345,10 +1273,10 @@ lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute \ \Union is an AC-operator\ -lemma Un_absorb1: "A \ B ==> A \ B = B" +lemma Un_absorb1: "A \ B \ A \ B = B" by (fact sup_absorb2) -lemma Un_absorb2: "B \ A ==> A \ B = A" +lemma Un_absorb2: "B \ A \ A \ B = A" by (fact sup_absorb1) lemma Un_empty_left: "{} \ B = B" @@ -1369,28 +1297,22 @@ lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" by blast -lemma Int_insert_left: - "(insert a B) Int C = (if a \ C then insert a (B \ C) else B \ C)" +lemma Int_insert_left: "(insert a B) \ C = (if a \ C then insert a (B \ C) else B \ C)" by auto -lemma Int_insert_left_if0[simp]: - "a \ C \ (insert a B) Int C = B \ C" +lemma Int_insert_left_if0 [simp]: "a \ C \ (insert a B) \ C = B \ C" by auto -lemma Int_insert_left_if1[simp]: - "a \ C \ (insert a B) Int C = insert a (B Int C)" +lemma Int_insert_left_if1 [simp]: "a \ C \ (insert a B) \ C = insert a (B \ C)" by auto -lemma Int_insert_right: - "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" +lemma Int_insert_right: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" by auto -lemma Int_insert_right_if0[simp]: - "a \ A \ A Int (insert a B) = A Int B" +lemma Int_insert_right_if0 [simp]: "a \ A \ A \ (insert a B) = A \ B" by auto -lemma Int_insert_right_if1[simp]: - "a \ A \ A Int (insert a B) = insert a (A Int B)" +lemma Int_insert_right_if1 [simp]: "a \ A \ A \ (insert a B) = insert a (A \ B)" by auto lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" @@ -1399,17 +1321,16 @@ lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact sup_inf_distrib2) -lemma Un_Int_crazy: - "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" +lemma Un_Int_crazy: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" by blast -lemma subset_Un_eq: "(A \ B) = (A \ B = B)" +lemma subset_Un_eq: "A \ B \ A \ B = B" by (fact le_iff_sup) -lemma Un_empty [iff]: "(A \ B = {}) = (A = {} & B = {})" +lemma Un_empty [iff]: "A \ B = {} \ A = {} \ B = {}" by (fact sup_eq_bot_iff) (* FIXME: already simp *) -lemma Un_subset_iff [simp]: "(A \ B \ C) = (A \ C & B \ C)" +lemma Un_subset_iff [simp]: "A \ B \ C \ A \ C \ B \ C" by (fact le_sup_iff) lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" @@ -1419,78 +1340,79 @@ by blast -text \\medskip Set complement\ - -lemma Compl_disjoint [simp]: "A \ -A = {}" +text \\<^medskip> Set complement\ + +lemma Compl_disjoint [simp]: "A \ - A = {}" by (fact inf_compl_bot) -lemma Compl_disjoint2 [simp]: "-A \ A = {}" +lemma Compl_disjoint2 [simp]: "- A \ A = {}" by (fact compl_inf_bot) -lemma Compl_partition: "A \ -A = UNIV" +lemma Compl_partition: "A \ - A = UNIV" by (fact sup_compl_top) -lemma Compl_partition2: "-A \ A = UNIV" +lemma Compl_partition2: "- A \ A = UNIV" by (fact compl_sup_top) -lemma double_complement: "- (-A) = (A::'a set)" +lemma double_complement: "- (-A) = A" for A :: "'a set" by (fact double_compl) (* already simp *) -lemma Compl_Un: "-(A \ B) = (-A) \ (-B)" +lemma Compl_Un: "- (A \ B) = (- A) \ (- B)" by (fact compl_sup) (* already simp *) -lemma Compl_Int: "-(A \ B) = (-A) \ (-B)" +lemma Compl_Int: "- (A \ B) = (- A) \ (- B)" by (fact compl_inf) (* already simp *) -lemma subset_Compl_self_eq: "(A \ -A) = (A = {})" +lemma subset_Compl_self_eq: "A \ - A \ A = {}" by blast -lemma Un_Int_assoc_eq: "((A \ B) \ C = A \ (B \ C)) = (C \ A)" +lemma Un_Int_assoc_eq: "(A \ B) \ C = A \ (B \ C) \ C \ A" \ \Halmos, Naive Set Theory, page 16.\ by blast -lemma Compl_UNIV_eq: "-UNIV = {}" +lemma Compl_UNIV_eq: "- UNIV = {}" by (fact compl_top_eq) (* already simp *) -lemma Compl_empty_eq: "-{} = UNIV" +lemma Compl_empty_eq: "- {} = UNIV" by (fact compl_bot_eq) (* already simp *) -lemma Compl_subset_Compl_iff [iff]: "(-A \ -B) = (B \ A)" +lemma Compl_subset_Compl_iff [iff]: "- A \ - B \ B \ A" by (fact compl_le_compl_iff) (* FIXME: already simp *) -lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))" +lemma Compl_eq_Compl_iff [iff]: "- A = - B \ A = B" for A B :: "'a set" by (fact compl_eq_compl_iff) (* FIXME: already simp *) -lemma Compl_insert: "- insert x A = (-A) - {x}" +lemma Compl_insert: "- insert x A = (- A) - {x}" by blast -text \\medskip Bounded quantifiers. +text \\<^medskip> Bounded quantifiers. The following are not added to the default simpset because - (a) they duplicate the body and (b) there are no similar rules for \Int\.\ - -lemma ball_Un: "(\x \ A \ B. P x) = ((\x\A. P x) & (\x\B. P x))" + (a) they duplicate the body and (b) there are no similar rules for \Int\. +\ + +lemma ball_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast -lemma bex_Un: "(\x \ A \ B. P x) = ((\x\A. P x) | (\x\B. P x))" +lemma bex_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast -text \\medskip Set difference.\ - -lemma Diff_eq: "A - B = A \ (-B)" +text \\<^medskip> Set difference.\ + +lemma Diff_eq: "A - B = A \ (- B)" by blast -lemma Diff_eq_empty_iff [simp]: "(A - B = {}) = (A \ B)" +lemma Diff_eq_empty_iff [simp]: "A - B = {} \ A \ B" by blast lemma Diff_cancel [simp]: "A - A = {}" by blast -lemma Diff_idemp [simp]: "(A - B) - B = A - (B::'a set)" -by blast - -lemma Diff_triv: "A \ B = {} ==> A - B = A" +lemma Diff_idemp [simp]: "(A - B) - B = A - B" for A B :: "'a set" + by blast + +lemma Diff_triv: "A \ B = {} \ A - B = A" by (blast elim: equalityE) lemma empty_Diff [simp]: "{} - A = {}" @@ -1502,39 +1424,39 @@ lemma Diff_UNIV [simp]: "A - UNIV = {}" by blast -lemma Diff_insert0 [simp]: "x \ A ==> A - insert x B = A - B" +lemma Diff_insert0 [simp]: "x \ A \ A - insert x B = A - B" by blast lemma Diff_insert: "A - insert a B = A - B - {a}" - \ \NOT SUITABLE FOR REWRITING since \{a} == insert a 0\\ + \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma Diff_insert2: "A - insert a B = A - {a} - B" - \ \NOT SUITABLE FOR REWRITING since \{a} == insert a 0\\ + \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" by auto -lemma insert_Diff1 [simp]: "x \ B ==> insert x A - B = A - B" +lemma insert_Diff1 [simp]: "x \ B \ insert x A - B = A - B" by blast lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" -by blast - -lemma insert_Diff: "a \ A ==> insert a (A - {a}) = A" by blast -lemma Diff_insert_absorb: "x \ A ==> (insert x A) - {x} = A" +lemma insert_Diff: "a \ A \ insert a (A - {a}) = A" + by blast + +lemma Diff_insert_absorb: "x \ A \ (insert x A) - {x} = A" by auto lemma Diff_disjoint [simp]: "A \ (B - A) = {}" by blast -lemma Diff_partition: "A \ B ==> A \ (B - A) = B" +lemma Diff_partition: "A \ B \ A \ (B - A) = B" by blast -lemma double_diff: "A \ B ==> B \ C ==> B - (C - A) = A" +lemma double_diff: "A \ B \ B \ C \ B - (C - A) = A" by blast lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" @@ -1567,13 +1489,13 @@ lemma Diff_Compl [simp]: "A - (- B) = A \ B" by auto -lemma Compl_Diff_eq [simp]: "- (A - B) = -A \ B" +lemma Compl_Diff_eq [simp]: "- (A - B) = - A \ B" by blast -lemma subset_Compl_singleton [simp]: "A \ - {b} \ (b \ A)" +lemma subset_Compl_singleton [simp]: "A \ - {b} \ b \ A" by blast -text \\medskip Quantification over type @{typ bool}.\ +text \\<^medskip> Quantification over type @{typ bool}.\ lemma bool_induct: "P True \ P False \ P x" by (cases x) auto @@ -1590,7 +1512,7 @@ lemma UNIV_bool: "UNIV = {False, True}" by (auto intro: bool_induct) -text \\medskip \Pow\\ +text \\<^medskip> \Pow\\ lemma Pow_empty [simp]: "Pow {} = {{}}" by (auto simp add: Pow_def) @@ -1601,7 +1523,7 @@ lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" by (blast intro: image_eqI [where ?x = "u - {a}" for u]) -lemma Pow_Compl: "Pow (- A) = {-B | B. A \ Pow B}" +lemma Pow_Compl: "Pow (- A) = {- B | B. A \ Pow B}" by (blast intro: exI [where ?x = "- u" for u]) lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" @@ -1614,21 +1536,21 @@ by blast -text \\medskip Miscellany.\ - -lemma set_eq_subset: "(A = B) = (A \ B & B \ A)" +text \\<^medskip> Miscellany.\ + +lemma set_eq_subset: "A = B \ A \ B \ B \ A" by blast -lemma subset_iff: "(A \ B) = (\t. t \ A --> t \ B)" +lemma subset_iff: "A \ B \ (\t. t \ A \ t \ B)" by blast -lemma subset_iff_psubset_eq: "(A \ B) = ((A \ B) | (A = B))" - by (unfold less_le) blast - -lemma all_not_in_conv [simp]: "(\x. x \ A) = (A = {})" +lemma subset_iff_psubset_eq: "A \ B \ A \ B \ A = B" + unfolding less_le by blast + +lemma all_not_in_conv [simp]: "(\x. x \ A) \ A = {}" by blast -lemma ex_in_conv: "(\x. x \ A) = (A \ {})" +lemma ex_in_conv: "(\x. x \ A) \ A \ {}" by blast lemma ball_simps [simp, no_atp]: @@ -1658,112 +1580,110 @@ subsubsection \Monotonicity of various operations\ -lemma image_mono: "A \ B ==> f`A \ f`B" +lemma image_mono: "A \ B \ f ` A \ f ` B" by blast -lemma Pow_mono: "A \ B ==> Pow A \ Pow B" +lemma Pow_mono: "A \ B \ Pow A \ Pow B" by blast -lemma insert_mono: "C \ D ==> insert a C \ insert a D" +lemma insert_mono: "C \ D \ insert a C \ insert a D" by blast -lemma Un_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" +lemma Un_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact sup_mono) -lemma Int_mono: "A \ C ==> B \ D ==> A \ B \ C \ D" +lemma Int_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact inf_mono) -lemma Diff_mono: "A \ C ==> D \ B ==> A - B \ C - D" +lemma Diff_mono: "A \ C \ D \ B \ A - B \ C - D" by blast -lemma Compl_anti_mono: "A \ B ==> -B \ -A" +lemma Compl_anti_mono: "A \ B \ - B \ - A" by (fact compl_mono) -text \\medskip Monotonicity of implications.\ - -lemma in_mono: "A \ B ==> x \ A --> x \ B" +text \\<^medskip> Monotonicity of implications.\ + +lemma in_mono: "A \ B \ x \ A \ x \ B" apply (rule impI) apply (erule subsetD, assumption) done -lemma conj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 & P2) --> (Q1 & Q2)" - by iprover - -lemma disj_mono: "P1 --> Q1 ==> P2 --> Q2 ==> (P1 | P2) --> (Q1 | Q2)" +lemma conj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover -lemma imp_mono: "Q1 --> P1 ==> P2 --> Q2 ==> (P1 --> P2) --> (Q1 --> Q2)" +lemma disj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover -lemma imp_refl: "P --> P" .. - -lemma not_mono: "Q --> P ==> ~ P --> ~ Q" +lemma imp_mono: "Q1 \ P1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover -lemma ex_mono: "(!!x. P x --> Q x) ==> (EX x. P x) --> (EX x. Q x)" +lemma imp_refl: "P \ P" .. + +lemma not_mono: "Q \ P \ \ P \ \ Q" by iprover -lemma all_mono: "(!!x. P x --> Q x) ==> (ALL x. P x) --> (ALL x. Q x)" +lemma ex_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover -lemma Collect_mono: "(!!x. P x --> Q x) ==> Collect P \ Collect Q" +lemma all_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" + by iprover + +lemma Collect_mono: "(\x. P x \ Q x) \ Collect P \ Collect Q" by blast -lemma Int_Collect_mono: - "A \ B ==> (!!x. x \ A ==> P x --> Q x) ==> A \ Collect P \ B \ Collect Q" +lemma Int_Collect_mono: "A \ B \ (\x. x \ A \ P x \ Q x) \ A \ Collect P \ B \ Collect Q" by blast lemmas basic_monos = - subset_refl imp_refl disj_mono conj_mono - ex_mono Collect_mono in_mono - -lemma eq_to_mono: "a = b ==> c = d ==> b --> d ==> a --> c" + subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono in_mono + +lemma eq_to_mono: "a = b \ c = d \ b \ d \ a \ c" by iprover subsubsection \Inverse image of a function\ -definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where - "f -` B == {x. f x : B}" - -lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" - by (unfold vimage_def) blast - -lemma vimage_singleton_eq: "(a : f -` {b}) = (f a = b)" +definition vimage :: "('a \ 'b) \ 'b set \ 'a set" (infixr "-`" 90) + where "f -` B \ {x. f x \ B}" + +lemma vimage_eq [simp]: "a \ f -` B \ f a \ B" + unfolding vimage_def by blast + +lemma vimage_singleton_eq: "a \ f -` {b} \ f a = b" by simp -lemma vimageI [intro]: "f a = b ==> b:B ==> a : f -` B" - by (unfold vimage_def) blast - -lemma vimageI2: "f a : A ==> a : f -` A" - by (unfold vimage_def) fast - -lemma vimageE [elim!]: "a: f -` B ==> (!!x. f a = x ==> x:B ==> P) ==> P" - by (unfold vimage_def) blast - -lemma vimageD: "a : f -` A ==> f a : A" - by (unfold vimage_def) fast +lemma vimageI [intro]: "f a = b \ b \ B \ a \ f -` B" + unfolding vimage_def by blast + +lemma vimageI2: "f a \ A \ a \ f -` A" + unfolding vimage_def by fast + +lemma vimageE [elim!]: "a \ f -` B \ (\x. f a = x \ x \ B \ P) \ P" + unfolding vimage_def by blast + +lemma vimageD: "a \ f -` A \ f a \ A" + unfolding vimage_def by fast lemma vimage_empty [simp]: "f -` {} = {}" by blast -lemma vimage_Compl: "f -` (-A) = -(f -` A)" +lemma vimage_Compl: "f -` (- A) = - (f -` A)" by blast -lemma vimage_Un [simp]: "f -` (A Un B) = (f -` A) Un (f -` B)" +lemma vimage_Un [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by blast -lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)" +lemma vimage_Int [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by fast lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" by blast -lemma vimage_Collect: "(!!x. P (f x) = Q x) ==> f -` (Collect P) = Collect Q" +lemma vimage_Collect: "(\x. P (f x) = Q x) \ f -` (Collect P) = Collect Q" by blast -lemma vimage_insert: "f-`(insert a B) = (f-`{a}) Un (f-`B)" - \ \NOT suitable for rewriting because of the recurrence of @{term "{a}"}.\ +lemma vimage_insert: "f -` (insert a B) = (f -` {a}) \ (f -` B)" + \ \NOT suitable for rewriting because of the recurrence of \{a}\.\ by blast lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" @@ -1772,18 +1692,18 @@ lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" by blast -lemma vimage_mono: "A \ B ==> f -` A \ f -` B" +lemma vimage_mono: "A \ B \ f -` A \ f -` B" \ \monotonicity\ by blast -lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}" -by (blast intro: sym) - -lemma image_vimage_subset: "f ` (f -` A) <= A" -by blast - -lemma image_vimage_eq [simp]: "f ` (f -` A) = A Int range f" -by blast +lemma vimage_image_eq: "f -` (f ` A) = {y. \x\A. f x = f y}" + by (blast intro: sym) + +lemma image_vimage_subset: "f ` (f -` A) \ A" + by blast + +lemma image_vimage_eq [simp]: "f ` (f -` A) = A \ range f" + by blast lemma image_subset_iff_subset_vimage: "f ` A \ B \ A \ f -` B" by blast @@ -1793,21 +1713,20 @@ lemma vimage_if [simp]: "((\x. if x \ B then c else d) -` A) = (if c \ A then (if d \ A then UNIV else B) - else if d \ A then -B else {})" + else if d \ A then - B else {})" by (auto simp add: vimage_def) -lemma vimage_inter_cong: - "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" +lemma vimage_inter_cong: "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" by auto -lemma vimage_ident [simp]: "(%x. x) -` Y = Y" +lemma vimage_ident [simp]: "(\x. x) -` Y = Y" by blast subsubsection \Singleton sets\ -definition is_singleton :: "'a set \ bool" where - "is_singleton A \ (\x. A = {x})" +definition is_singleton :: "'a set \ bool" + where "is_singleton A \ (\x. A = {x})" lemma is_singletonI [simp, intro!]: "is_singleton {x}" unfolding is_singleton_def by simp @@ -1819,10 +1738,10 @@ unfolding is_singleton_def by blast -subsubsection \Getting the Contents of a Singleton Set\ - -definition the_elem :: "'a set \ 'a" where - "the_elem X = (THE x. X = {x})" +subsubsection \Getting the contents of a singleton set\ + +definition the_elem :: "'a set \ 'a" + where "the_elem X = (THE x. X = {x})" lemma the_elem_eq [simp]: "the_elem {x} = x" by (simp add: the_elem_def) @@ -1832,9 +1751,10 @@ lemma the_elem_image_unique: assumes "A \ {}" - assumes *: "\y. y \ A \ f y = f x" + and *: "\y. y \ A \ f y = f x" shows "the_elem (f ` A) = f x" -unfolding the_elem_def proof (rule the1_equality) + unfolding the_elem_def +proof (rule the1_equality) from \A \ {}\ obtain y where "y \ A" by auto with * have "f x = f y" by simp with \y \ A\ have "f x \ f ` A" by blast @@ -1845,12 +1765,12 @@ subsubsection \Least value operator\ -lemma Least_mono: - "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y - ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)" - \ \Courtesy of Stephan Merz\ +lemma Least_mono: "mono f \ \x\S. \y\S. x \ y \ (LEAST y. y \ f ` S) = f (LEAST x. x \ S)" + for f :: "'a::order \ 'b::order" + \ \Courtesy of Stephan Merz\ apply clarify - apply (erule_tac P = "%x. x : S" in LeastI2_order, fast) + apply (erule_tac P = "\x. x : S" in LeastI2_order) + apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) done @@ -1858,22 +1778,18 @@ subsubsection \Monad operation\ -definition bind :: "'a set \ ('a \ 'b set) \ 'b set" where - "bind A f = {x. \B \ f`A. x \ B}" +definition bind :: "'a set \ ('a \ 'b set) \ 'b set" + where "bind A f = {x. \B \ f`A. x \ B}" hide_const (open) bind -lemma bind_bind: - fixes A :: "'a set" - shows "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" +lemma bind_bind: "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" for A :: "'a set" by (auto simp add: bind_def) -lemma empty_bind [simp]: - "Set.bind {} f = {}" +lemma empty_bind [simp]: "Set.bind {} f = {}" by (simp add: bind_def) -lemma nonempty_bind_const: - "A \ {} \ Set.bind A (\_. B) = B" +lemma nonempty_bind_const: "A \ {} \ Set.bind A (\_. B) = B" by (auto simp add: bind_def) lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" @@ -1882,53 +1798,50 @@ lemma bind_singleton_conv_image: "Set.bind A (\x. {f x}) = f ` A" by(auto simp add: bind_def) + subsubsection \Operations for execution\ -definition is_empty :: "'a set \ bool" where - [code_abbrev]: "is_empty A \ A = {}" +definition is_empty :: "'a set \ bool" + where [code_abbrev]: "is_empty A \ A = {}" hide_const (open) is_empty -definition remove :: "'a \ 'a set \ 'a set" where - [code_abbrev]: "remove x A = A - {x}" +definition remove :: "'a \ 'a set \ 'a set" + where [code_abbrev]: "remove x A = A - {x}" hide_const (open) remove -lemma member_remove [simp]: - "x \ Set.remove y A \ x \ A \ x \ y" +lemma member_remove [simp]: "x \ Set.remove y A \ x \ A \ x \ y" by (simp add: remove_def) -definition filter :: "('a \ bool) \ 'a set \ 'a set" where - [code_abbrev]: "filter P A = {a \ A. P a}" +definition filter :: "('a \ bool) \ 'a set \ 'a set" + where [code_abbrev]: "filter P A = {a \ A. P a}" hide_const (open) filter -lemma member_filter [simp]: - "x \ Set.filter P A \ x \ A \ P x" +lemma member_filter [simp]: "x \ Set.filter P A \ x \ A \ P x" by (simp add: filter_def) instantiation set :: (equal) equal begin -definition - "HOL.equal A B \ A \ B \ B \ A" - -instance proof -qed (auto simp add: equal_set_def) +definition "HOL.equal A B \ A \ B \ B \ A" + +instance by standard (auto simp add: equal_set_def) end text \Misc\ -definition "pairwise R S \ (\x \ S. \y\ S. x\y \ R x y)" - -lemma pairwise_subset: "\pairwise P S; T \ S\ \ pairwise P T" +definition "pairwise R S \ (\x \ S. \y\ S. x \ y \ R x y)" + +lemma pairwise_subset: "pairwise P S \ T \ S \ pairwise P T" by (force simp: pairwise_def) -definition disjnt where "disjnt A B \ A \ B = {}" - -lemma disjnt_iff: "disjnt A B \ (\x. ~ (x \ A \ x \ B))" +definition disjnt where "disjnt A B \ A \ B = {}" + +lemma disjnt_iff: "disjnt A B \ (\x. \ (x \ A \ x \ B))" by (force simp: disjnt_def) lemma pairwise_empty [simp]: "pairwise P {}" @@ -1938,12 +1851,11 @@ by (simp add: pairwise_def) lemma pairwise_insert: - "pairwise r (insert x s) \ (\y. y \ s \ y \ x \ r x y \ r y x) \ pairwise r s" -by (force simp: pairwise_def) - -lemma pairwise_image: - "pairwise r (f ` s) \ pairwise (\x y. (f x \ f y) \ r (f x) (f y)) s" -by (force simp: pairwise_def) + "pairwise r (insert x s) \ (\y. y \ s \ y \ x \ r x y \ r y x) \ pairwise r s" + by (force simp: pairwise_def) + +lemma pairwise_image: "pairwise r (f ` s) \ pairwise (\x y. (f x \ f y) \ r (f x) (f y)) s" + by (force simp: pairwise_def) lemma Int_emptyI: "(\x. x \ A \ x \ B \ False) \ A \ B = {}" by blast