diff -r 881e8e2cfec2 -r d0e2bad67bd4 src/HOL/Set.thy --- a/src/HOL/Set.thy Tue Aug 02 21:04:52 2016 +0200 +++ b/src/HOL/Set.thy Tue Aug 02 21:05:34 2016 +0200 @@ -1,9 +1,13 @@ -(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) +(* Title: HOL/Set.thy + Author: Tobias Nipkow + Author: Lawrence C Paulson + Author: Markus Wenzel +*) section \Set theory for higher-order logic\ theory Set -imports Lattices + imports Lattices begin subsection \Sets as predicates\ @@ -12,8 +16,8 @@ axiomatization Collect :: "('a \ bool) \ 'a set" \ "comprehension" and member :: "'a \ 'a set \ bool" \ "membership" -where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" - and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" + where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" + and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" notation member ("op \") and @@ -76,7 +80,8 @@ assumes "\x. x \ A \ x \ B" shows "A = B" proof - - from assms have "{x. x \ A} = {x. x \ B}" by simp + from assms have "{x. x \ A} = {x. x \ B}" + by simp then show ?thesis by simp qed @@ -347,7 +352,6 @@ by (simp add: Ball_def) text \Gives better instantiation for bound:\ - setup \ map_theory_claset (fn ctxt => ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) @@ -459,7 +463,7 @@ \ \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\.\ + \ \The same, with reversed premises for use with @{method erule} -- cf. @{thm rev_mp}.\ by (rule subsetD) lemma subsetCE [elim]: "A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" @@ -696,8 +700,7 @@ lemma UnI2 [elim?]: "c \ B \ c \ A \ B" by simp -text \\<^medskip> Classical introduction rule: no commitment to @{prop A} vs @{prop B}.\ - +text \\<^medskip> Classical introduction rule: no commitment to \A\ vs. \B\.\ lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by auto @@ -960,7 +963,7 @@ text \\<^medskip> Range of a function -- just an abbreviation for image!\ -abbreviation range :: "('a \ 'b) \ 'b set" \ "of function" +abbreviation range :: "('a \ 'b) \ 'b set" \ \of function\ where "range f \ f ` UNIV" lemma range_eqI: "b = f x \ b \ range f" @@ -1387,7 +1390,8 @@ 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" for 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}" @@ -1417,7 +1421,8 @@ lemma Diff_cancel [simp]: "A - A = {}" by blast -lemma Diff_idemp [simp]: "(A - B) - B = A - B" for A B :: "'a set" +lemma Diff_idemp [simp]: "(A - B) - B = A - B" + for A B :: "'a set" by blast lemma Diff_triv: "A \ B = {} \ A - B = A" @@ -1526,7 +1531,7 @@ by (auto simp add: Pow_def) lemma Pow_singleton_iff [simp]: "Pow X = {Y} \ X = {} \ Y = {}" - by blast + by blast (* somewhat slow *) lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" by (blast intro: image_eqI [where ?x = "u - {a}" for u]) @@ -1612,9 +1617,7 @@ text \\<^medskip> Monotonicity of implications.\ lemma in_mono: "A \ B \ x \ A \ x \ B" - apply (rule impI) - apply (erule subsetD, assumption) - done + by (rule impI) (erule subsetD) lemma conj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover @@ -1730,7 +1733,7 @@ lemma vimage_ident [simp]: "(\x. x) -` Y = Y" by blast - + subsubsection \Singleton sets\ definition is_singleton :: "'a set \ bool" @@ -1778,9 +1781,9 @@ \ \Courtesy of Stephan Merz\ apply clarify apply (erule_tac P = "\x. x : S" in LeastI2_order) - apply fast + apply fast apply (rule LeastI2_order) - apply (auto elim: monoD intro!: order_antisym) + apply (auto elim: monoD intro!: order_antisym) done @@ -1791,20 +1794,21 @@ hide_const (open) bind -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 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: bind_def) lemma empty_bind [simp]: "Set.bind {} f = {}" by (simp add: bind_def) lemma nonempty_bind_const: "A \ {} \ Set.bind A (\_. B) = B" - by (auto simp add: bind_def) + by (auto simp: bind_def) lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" - by (auto simp add: bind_def) + by (auto simp: bind_def) lemma bind_singleton_conv_image: "Set.bind A (\x. {f x}) = f ` A" - by(auto simp add: bind_def) + by (auto simp: bind_def) subsubsection \Operations for execution\ @@ -1842,12 +1846,14 @@ text \Misc\ -definition "pairwise R S \ (\x \ S. \y\ S. x \ y \ R x y)" +definition pairwise :: "('a \ 'a \ bool) \ 'a set \ bool" + where "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 = {}" +definition disjnt :: "'a set \ 'a set \ bool" + where "disjnt A B \ A \ B = {}" lemma disjnt_iff: "disjnt A B \ (\x. \ (x \ A \ x \ B))" by (force simp: disjnt_def)