# HG changeset patch # User wenzelm # Date 1466435030 -7200 # Node ID bc1f17d45e9127ca66c704fda4e8531d752df135 # Parent dff40165618cdb405c952a53a51ffac8a0caa17e misc tuning and modernization; diff -r dff40165618c -r bc1f17d45e91 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Sun Jun 19 22:51:42 2016 +0200 +++ b/src/HOL/Fun.thy Mon Jun 20 17:03:50 2016 +0200 @@ -11,11 +11,10 @@ keywords "functor" :: thy_goal begin -lemma apply_inverse: - "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" +lemma apply_inverse: "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" by auto -text\Uniqueness, so NOT the axiom of choice.\ +text \Uniqueness, so NOT the axiom of choice.\ lemma uniq_choice: "\x. \!y. Q x y \ \f. \x. Q x (f x)" by (force intro: theI') @@ -24,8 +23,8 @@ subsection \The Identity Function \id\\ -definition id :: "'a \ 'a" where - "id = (\x. x)" +definition id :: "'a \ 'a" + where "id = (\x. x)" lemma id_apply [simp]: "id x = x" by (simp add: id_def) @@ -51,55 +50,51 @@ notation (ASCII) comp (infixl "o" 55) -lemma comp_apply [simp]: "(f o g) x = f (g x)" +lemma comp_apply [simp]: "(f \ g) x = f (g x)" by (simp add: comp_def) -lemma comp_assoc: "(f o g) o h = f o (g o h)" +lemma comp_assoc: "(f \ g) \ h = f \ (g \ h)" by (simp add: fun_eq_iff) -lemma id_comp [simp]: "id o g = g" +lemma id_comp [simp]: "id \ g = g" by (simp add: fun_eq_iff) -lemma comp_id [simp]: "f o id = f" +lemma comp_id [simp]: "f \ id = f" by (simp add: fun_eq_iff) lemma comp_eq_dest: - "a o b = c o d \ a (b v) = c (d v)" + "a \ b = c \ d \ a (b v) = c (d v)" by (simp add: fun_eq_iff) lemma comp_eq_elim: - "a o b = c o d \ ((\v. a (b v) = c (d v)) \ R) \ R" + "a \ b = c \ d \ ((\v. a (b v) = c (d v)) \ R) \ R" by (simp add: fun_eq_iff) -lemma comp_eq_dest_lhs: "a o b = c \ a (b v) = c v" - by clarsimp - -lemma comp_eq_id_dest: "a o b = id o c \ a (b v) = c v" +lemma comp_eq_dest_lhs: "a \ b = c \ a (b v) = c v" by clarsimp -lemma image_comp: - "f ` (g ` r) = (f o g) ` r" +lemma comp_eq_id_dest: "a \ b = id \ c \ a (b v) = c v" + by clarsimp + +lemma image_comp: "f ` (g ` r) = (f \ g) ` r" by auto -lemma vimage_comp: - "f -` (g -` x) = (g \ f) -` x" +lemma vimage_comp: "f -` (g -` x) = (g \ f) -` x" by auto -lemma image_eq_imp_comp: "f ` A = g ` B \ (h o f) ` A = (h o g) ` B" +lemma image_eq_imp_comp: "f ` A = g ` B \ (h \ f) ` A = (h \ g) ` B" by (auto simp: comp_def elim!: equalityE) lemma image_bind: "f ` (Set.bind A g) = Set.bind A (op ` f \ g)" -by(auto simp add: Set.bind_def) + by (auto simp add: Set.bind_def) lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \ f)" -by(auto simp add: Set.bind_def) + by (auto simp add: Set.bind_def) -lemma (in group_add) minus_comp_minus [simp]: - "uminus \ uminus = id" +lemma (in group_add) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) -lemma (in boolean_algebra) minus_comp_minus [simp]: - "uminus \ uminus = id" +lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) code_printing @@ -108,8 +103,8 @@ subsection \The Forward Composition Operator \fcomp\\ -definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) where - "f \> g = (\x. g (f x))" +definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) + where "f \> g = (\x. g (f x))" lemma fcomp_apply [simp]: "(f \> g) x = g (f x)" by (simp add: fcomp_def) @@ -123,7 +118,7 @@ lemma fcomp_id [simp]: "f \> id = f" by (simp add: fcomp_def) -lemma fcomp_comp: "fcomp f g = comp g f" +lemma fcomp_comp: "fcomp f g = comp g f" by (simp add: ext) code_printing @@ -134,168 +129,143 @@ subsection \Mapping functions\ -definition map_fun :: "('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" where - "map_fun f g h = g \ h \ f" +definition map_fun :: "('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" + where "map_fun f g h = g \ h \ f" -lemma map_fun_apply [simp]: - "map_fun f g h x = g (h (f x))" +lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" by (simp add: map_fun_def) subsection \Injectivity and Bijectivity\ -definition inj_on :: "('a \ 'b) \ 'a set \ bool" where \ "injective" - "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" +definition inj_on :: "('a \ 'b) \ 'a set \ bool" \ \injective\ + where "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" -definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" where \ "bijective" - "bij_betw f A B \ inj_on f A \ f ` A = B" +definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" \ \bijective\ + where "bij_betw f A B \ inj_on f A \ f ` A = B" -text\A common special case: functions injective, surjective or bijective over -the entire domain type.\ +text \A common special case: functions injective, surjective or bijective over + the entire domain type.\ -abbreviation - "inj f \ inj_on f UNIV" +abbreviation "inj f \ inj_on f UNIV" -abbreviation surj :: "('a \ 'b) \ bool" where \ "surjective" - "surj f \ (range f = UNIV)" +abbreviation surj :: "('a \ 'b) \ bool" \ "surjective" + where "surj f \ range f = UNIV" -abbreviation - "bij f \ bij_betw f UNIV UNIV" +abbreviation "bij f \ bij_betw f UNIV UNIV" -text\The negated case:\ +text \The negated case:\ translations -"\ CONST surj f" <= "CONST range f \ CONST UNIV" - -lemma injI: - assumes "\x y. f x = f y \ x = y" - shows "inj f" - using assms unfolding inj_on_def by auto - -theorem range_ex1_eq: "inj f \ b : range f = (EX! x. b = f x)" - by (unfold inj_on_def, blast) + "\ CONST surj f" \ "CONST range f \ CONST UNIV" -lemma injD: "[| inj(f); f(x) = f(y) |] ==> x=y" -by (simp add: inj_on_def) - -lemma inj_on_eq_iff: "\inj_on f A; x \ A; y \ A\ \ (f x = f y) = (x = y)" -by (force simp add: inj_on_def) +lemma injI: "(\x y. f x = f y \ x = y) \ inj f" + unfolding inj_on_def by auto -lemma inj_on_cong: - "(\ a. a : A \ f a = g a) \ inj_on f A = inj_on g A" -unfolding inj_on_def by auto - -lemma inj_on_strict_subset: - "inj_on f B \ A \ B \ f ` A \ f ` B" +theorem range_ex1_eq: "inj f \ b \ range f \ (\!x. b = f x)" unfolding inj_on_def by blast -lemma inj_comp: - "inj f \ inj g \ inj (f \ g)" +lemma injD: "inj f \ f x = f y \ x = y" + by (simp add: inj_on_def) + +lemma inj_on_eq_iff: "inj_on f A \ x \ A \ y \ A \ f x = f y \ x = y" + by (force simp add: inj_on_def) + +lemma inj_on_cong: "(\a. a \ A \ f a = g a) \ inj_on f A = inj_on g A" + unfolding inj_on_def by auto + +lemma inj_on_strict_subset: "inj_on f B \ A \ B \ f ` A \ f ` B" + unfolding inj_on_def by blast + +lemma inj_comp: "inj f \ inj g \ inj (f \ g)" by (simp add: inj_on_def) lemma inj_fun: "inj f \ inj (\x y. f x)" by (simp add: inj_on_def fun_eq_iff) -lemma inj_eq: "inj f ==> (f(x) = f(y)) = (x=y)" -by (simp add: inj_on_eq_iff) +lemma inj_eq: "inj f \ f x = f y \ x = y" + by (simp add: inj_on_eq_iff) lemma inj_on_id[simp]: "inj_on id A" by (simp add: inj_on_def) -lemma inj_on_id2[simp]: "inj_on (%x. x) A" -by (simp add: inj_on_def) +lemma inj_on_id2[simp]: "inj_on (\x. x) A" + by (simp add: inj_on_def) lemma inj_on_Int: "inj_on f A \ inj_on f B \ inj_on f (A \ B)" -unfolding inj_on_def by blast + unfolding inj_on_def by blast lemma surj_id: "surj id" -by simp + by simp lemma bij_id[simp]: "bij id" -by (simp add: bij_betw_def) + by (simp add: bij_betw_def) -lemma bij_uminus: - fixes x :: "'a :: ab_group_add" - shows "bij (uminus :: 'a\'a)" -unfolding bij_betw_def inj_on_def -by (force intro: minus_minus [symmetric]) +lemma bij_uminus: "bij (uminus :: 'a \ 'a::ab_group_add)" + unfolding bij_betw_def inj_on_def + by (force intro: minus_minus [symmetric]) -lemma inj_onI [intro?]: - "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" -by (simp add: inj_on_def) +lemma inj_onI [intro?]: "(\x y. x \ A \ y \ A \ f x = f y \ x = y) \ inj_on f A" + by (simp add: inj_on_def) -lemma inj_on_inverseI: "(!!x. x:A ==> g(f(x)) = x) ==> inj_on f A" -by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) - -lemma inj_onD: "[| inj_on f A; f(x)=f(y); x:A; y:A |] ==> x=y" -by (unfold inj_on_def, blast) +lemma inj_on_inverseI: "(\x. x \ A \ g (f x) = x) \ inj_on f A" + by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) -lemma comp_inj_on: - "[| inj_on f A; inj_on g (f`A) |] ==> inj_on (g o f) A" -by (simp add: comp_def inj_on_def) +lemma inj_onD: "inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y" + unfolding inj_on_def by blast -lemma inj_on_imageI: "inj_on (g o f) A \ inj_on g (f ` A)" +lemma comp_inj_on: "inj_on f A \ inj_on g (f ` A) \ inj_on (g \ f) A" + by (simp add: comp_def inj_on_def) + +lemma inj_on_imageI: "inj_on (g \ f) A \ inj_on g (f ` A)" by (auto simp add: inj_on_def) -lemma inj_on_image_iff: "\ ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y); - inj_on f A \ \ inj_on g (f ` A) = inj_on g A" -unfolding inj_on_def by blast +lemma inj_on_image_iff: + "\x\A. \y\A. g (f x) = g (f y) \ g x = g y \ inj_on f A \ inj_on g (f ` A) = inj_on g A" + unfolding inj_on_def by blast -lemma inj_on_contraD: "[| inj_on f A; ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)" -unfolding inj_on_def by blast +lemma inj_on_contraD: "inj_on f A \ x \ y \ x \ A \ y \ A \ f x \ f y" + unfolding inj_on_def by blast lemma inj_singleton [simp]: "inj_on (\x. {x}) A" by (simp add: inj_on_def) lemma inj_on_empty[iff]: "inj_on f {}" -by(simp add: inj_on_def) - -lemma subset_inj_on: "[| inj_on f B; A <= B |] ==> inj_on f A" -unfolding inj_on_def by blast + by (simp add: inj_on_def) -lemma inj_on_Un: - "inj_on f (A Un B) = - (inj_on f A & inj_on f B & f`(A-B) Int f`(B-A) = {})" -apply(unfold inj_on_def) -apply (blast intro:sym) -done +lemma subset_inj_on: "inj_on f B \ A \ B \ inj_on f A" + unfolding inj_on_def by blast + +lemma inj_on_Un: "inj_on f (A \ B) \ inj_on f A \ inj_on f B \ f ` (A - B) \ f ` (B - A) = {}" + unfolding inj_on_def by (blast intro: sym) -lemma inj_on_insert[iff]: - "inj_on f (insert a A) = (inj_on f A & f a ~: f`(A-{a}))" -apply(unfold inj_on_def) -apply (blast intro:sym) -done +lemma inj_on_insert [iff]: "inj_on f (insert a A) \ inj_on f A \ f a \ f ` (A - {a})" + unfolding inj_on_def by (blast intro: sym) + +lemma inj_on_diff: "inj_on f A \ inj_on f (A - B)" + unfolding inj_on_def by blast -lemma inj_on_diff: "inj_on f A ==> inj_on f (A-B)" -apply(unfold inj_on_def) -apply (blast) -done +lemma comp_inj_on_iff: "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' \ f) A" + by (auto simp add: comp_inj_on inj_on_def) -lemma comp_inj_on_iff: - "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' o f) A" -by(auto simp add: comp_inj_on inj_on_def) - -lemma inj_on_imageI2: - "inj_on (f' o f) A \ inj_on f A" -by(auto simp add: comp_inj_on inj_on_def) +lemma inj_on_imageI2: "inj_on (f' \ f) A \ inj_on f A" + by (auto simp add: comp_inj_on inj_on_def) lemma inj_img_insertE: assumes "inj_on f A" - assumes "x \ B" and "insert x B = f ` A" - obtains x' A' where "x' \ A'" and "A = insert x' A'" - and "x = f x'" and "B = f ` A'" + assumes "x \ B" + and "insert x B = f ` A" + obtains x' A' where "x' \ A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" proof - from assms have "x \ f ` A" by auto then obtain x' where *: "x' \ A" "x = f x'" by auto - then have "A = insert x' (A - {x'})" by auto - with assms * have "B = f ` (A - {x'})" - by (auto dest: inj_on_contraD) + then have A: "A = insert x' (A - {x'})" by auto + with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD) have "x' \ A - {x'}" by simp - from \x' \ A - {x'}\ \A = insert x' (A - {x'})\ \x = f x'\ \B = image f (A - {x'})\ - show ?thesis .. + from this A \x = f x'\ B show ?thesis .. qed lemma linorder_injI: - assumes hyp: "\x y. x < (y::'a::linorder) \ f x \ f y" + assumes hyp: "\x y::'a::linorder. x < y \ f x \ f y" shows "inj f" \ \Courtesy of Stephan Merz\ proof (rule inj_onI) @@ -307,7 +277,9 @@ lemma surj_def: "surj f \ (\y. \x. y = f x)" by auto -lemma surjI: assumes *: "\ x. g (f x) = x" shows "surj g" +lemma surjI: + assumes *: "\ x. g (f x) = x" + shows "surj g" using *[symmetric] by auto lemma surjD: "surj f \ \x. y = f x" @@ -316,15 +288,17 @@ lemma surjE: "surj f \ (\x. y = f x \ C) \ C" by (simp add: surj_def, blast) -lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" -apply (simp add: comp_def surj_def, clarify) -apply (drule_tac x = y in spec, clarify) -apply (drule_tac x = x in spec, blast) -done +lemma comp_surj: "surj f \ surj g \ surj (g \ f)" + apply (simp add: comp_def surj_def) + apply clarify + apply (drule_tac x = y in spec) + apply clarify + apply (drule_tac x = x in spec) + apply blast + done -lemma bij_betw_imageI: - "\ inj_on f A; f ` A = B \ \ bij_betw f A B" -unfolding bij_betw_def by clarify +lemma bij_betw_imageI: "inj_on f A \ f ` A = B \ bij_betw f A B" + unfolding bij_betw_def by clarify lemma bij_betw_imp_surj_on: "bij_betw f A B \ f ` A = B" unfolding bij_betw_def by clarify @@ -332,122 +306,119 @@ lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" unfolding bij_betw_def by auto -lemma bij_betw_empty1: - assumes "bij_betw f {} A" - shows "A = {}" -using assms unfolding bij_betw_def by blast +lemma bij_betw_empty1: "bij_betw f {} A \ A = {}" + unfolding bij_betw_def by blast -lemma bij_betw_empty2: - assumes "bij_betw f A {}" - shows "A = {}" -using assms unfolding bij_betw_def by blast +lemma bij_betw_empty2: "bij_betw f A {} \ A = {}" + unfolding bij_betw_def by blast -lemma inj_on_imp_bij_betw: - "inj_on f A \ bij_betw f A (f ` A)" -unfolding bij_betw_def by simp +lemma inj_on_imp_bij_betw: "inj_on f A \ bij_betw f A (f ` A)" + unfolding bij_betw_def by simp lemma bij_def: "bij f \ inj f \ surj f" unfolding bij_betw_def .. -lemma bijI: "[| inj f; surj f |] ==> bij f" -by (simp add: bij_def) +lemma bijI: "inj f \ surj f \ bij f" + by (simp add: bij_def) -lemma bij_is_inj: "bij f ==> inj f" -by (simp add: bij_def) +lemma bij_is_inj: "bij f \ inj f" + by (simp add: bij_def) -lemma bij_is_surj: "bij f ==> surj f" -by (simp add: bij_def) +lemma bij_is_surj: "bij f \ surj f" + by (simp add: bij_def) lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" -by (simp add: bij_betw_def) + by (simp add: bij_betw_def) -lemma bij_betw_trans: - "bij_betw f A B \ bij_betw g B C \ bij_betw (g o f) A C" -by(auto simp add:bij_betw_def comp_inj_on) +lemma bij_betw_trans: "bij_betw f A B \ bij_betw g B C \ bij_betw (g \ f) A C" + by (auto simp add:bij_betw_def comp_inj_on) -lemma bij_comp: "bij f \ bij g \ bij (g o f)" +lemma bij_comp: "bij f \ bij g \ bij (g \ f)" by (rule bij_betw_trans) -lemma bij_betw_comp_iff: - "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' o f) A A''" -by(auto simp add: bij_betw_def inj_on_def) +lemma bij_betw_comp_iff: "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' \ f) A A''" + by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_comp_iff2: - assumes BIJ: "bij_betw f' A' A''" and IM: "f ` A \ A'" - shows "bij_betw f A A' \ bij_betw (f' o f) A A''" -using assms -proof(auto simp add: bij_betw_comp_iff) + assumes bij: "bij_betw f' A' A''" + and img: "f ` A \ A'" + shows "bij_betw f A A' \ bij_betw (f' \ f) A A''" + using assms +proof (auto simp add: bij_betw_comp_iff) assume *: "bij_betw (f' \ f) A A''" - thus "bij_betw f A A'" - using IM - proof(auto simp add: bij_betw_def) + then show "bij_betw f A A'" + using img + proof (auto simp add: bij_betw_def) assume "inj_on (f' \ f) A" - thus "inj_on f A" using inj_on_imageI2 by blast + then show "inj_on f A" using inj_on_imageI2 by blast next - fix a' assume **: "a' \ A'" - hence "f' a' \ A''" using BIJ unfolding bij_betw_def by auto - then obtain a where 1: "a \ A \ f'(f a) = f' a'" using * - unfolding bij_betw_def by force - hence "f a \ A'" using IM by auto - hence "f a = a'" using BIJ ** 1 unfolding bij_betw_def inj_on_def by auto - thus "a' \ f ` A" using 1 by auto + fix a' + assume **: "a' \ A'" + then have "f' a' \ A''" using bij unfolding bij_betw_def by auto + then obtain a where 1: "a \ A \ f'(f a) = f' a'" + using * unfolding bij_betw_def by force + then have "f a \ A'" using img by auto + then have "f a = a'" + using bij ** 1 unfolding bij_betw_def inj_on_def by auto + then show "a' \ f ` A" + using 1 by auto qed qed -lemma bij_betw_inv: assumes "bij_betw f A B" shows "EX g. bij_betw g B A" +lemma bij_betw_inv: + assumes "bij_betw f A B" + shows "\g. bij_betw g B A" proof - have i: "inj_on f A" and s: "f ` A = B" - using assms by(auto simp:bij_betw_def) - let ?P = "%b a. a:A \ f a = b" let ?g = "%b. The (?P b)" - { fix a b assume P: "?P b a" - hence ex1: "\a. ?P b a" using s by blast - hence uex1: "\!a. ?P b a" by(blast dest:inj_onD[OF i]) - hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp - } note g = this + using assms by (auto simp: bij_betw_def) + let ?P = "\b a. a \ A \ f a = b" + let ?g = "\b. The (?P b)" + have g: "?g b = a" if P: "?P b a" for a b + proof - + from that have ex1: "\a. ?P b a" using s by blast + then have uex1: "\!a. ?P b a" by (blast dest:inj_onD[OF i]) + then show ?thesis using the1_equality[OF uex1, OF P] P by simp + qed have "inj_on ?g B" - proof(rule inj_onI) - fix x y assume "x:B" "y:B" "?g x = ?g y" - from s \x:B\ obtain a1 where a1: "?P x a1" by blast - from s \y:B\ obtain a2 where a2: "?P y a2" by blast - from g[OF a1] a1 g[OF a2] a2 \?g x = ?g y\ show "x=y" by simp + proof (rule inj_onI) + fix x y + assume "x \ B" "y \ B" "?g x = ?g y" + from s \x \ B\ obtain a1 where a1: "?P x a1" by blast + from s \y \ B\ obtain a2 where a2: "?P y a2" by blast + from g [OF a1] a1 g [OF a2] a2 \?g x = ?g y\ show "x = y" by simp qed moreover have "?g ` B = A" - proof(auto simp: image_def) - fix b assume "b:B" + proof (auto simp: image_def) + fix b + assume "b \ B" with s obtain a where P: "?P b a" by blast - thus "?g b \ A" using g[OF P] by auto + then show "?g b \ A" using g[OF P] by auto next - fix a assume "a:A" + fix a + assume "a \ A" then obtain b where P: "?P b a" using s by blast - then have "b:B" using s by blast + then have "b \ B" using s by blast with g[OF P] show "\b\B. a = ?g b" by blast qed - ultimately show ?thesis by(auto simp:bij_betw_def) + ultimately show ?thesis by (auto simp: bij_betw_def) qed -lemma bij_betw_cong: - "(\ a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" -unfolding bij_betw_def inj_on_def by force +lemma bij_betw_cong: "(\ a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" + unfolding bij_betw_def inj_on_def by force -lemma bij_betw_id[intro, simp]: - "bij_betw id A A" -unfolding bij_betw_def id_def by auto +lemma bij_betw_id[intro, simp]: "bij_betw id A A" + unfolding bij_betw_def id_def by auto -lemma bij_betw_id_iff: - "bij_betw id A B \ A = B" -by(auto simp add: bij_betw_def) +lemma bij_betw_id_iff: "bij_betw id A B \ A = B" + by (auto simp add: bij_betw_def) lemma bij_betw_combine: assumes "bij_betw f A B" "bij_betw f C D" "B \ D = {}" shows "bij_betw f (A \ C) (B \ D)" using assms unfolding bij_betw_def inj_on_Un image_Un by auto -lemma bij_betw_subset: - assumes BIJ: "bij_betw f A A'" and - SUB: "B \ A" and IM: "f ` B = B'" - shows "bij_betw f B B'" -using assms -by(unfold bij_betw_def inj_on_def, auto simp add: inj_on_def) +lemma bij_betw_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw f B B'" + by (auto simp add: bij_betw_def inj_on_def) lemma bij_pointE: assumes "bij f" @@ -460,85 +431,77 @@ with that show thesis by blast qed -lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" -by simp +lemma surj_image_vimage_eq: "surj f \ f ` (f -` A) = A" + by simp lemma surj_vimage_empty: - assumes "surj f" shows "f -` A = {} \ A = {}" - using surj_image_vimage_eq[OF \surj f\, of A] + assumes "surj f" + shows "f -` A = {} \ A = {}" + using surj_image_vimage_eq [OF \surj f\, of A] by (intro iffI) fastforce+ -lemma inj_vimage_image_eq: "inj f ==> f -` (f ` A) = A" -by (simp add: inj_on_def, blast) +lemma inj_vimage_image_eq: "inj f \ f -` (f ` A) = A" + unfolding inj_on_def by blast -lemma vimage_subsetD: "surj f ==> f -` B <= A ==> B <= f ` A" -by (blast intro: sym) +lemma vimage_subsetD: "surj f \ f -` B \ A \ B \ f ` A" + by (blast intro: sym) -lemma vimage_subsetI: "inj f ==> B <= f ` A ==> f -` B <= A" -by (unfold inj_on_def, blast) +lemma vimage_subsetI: "inj f \ B \ f ` A \ f -` B \ A" + unfolding inj_on_def by blast -lemma vimage_subset_eq: "bij f ==> (f -` B <= A) = (B <= f ` A)" -apply (unfold bij_def) -apply (blast del: subsetI intro: vimage_subsetI vimage_subsetD) -done +lemma vimage_subset_eq: "bij f \ f -` B \ A \ B \ f ` A" + unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) -lemma inj_on_image_eq_iff: "\ inj_on f C; A \ C; B \ C \ \ f ` A = f ` B \ A = B" -by(fastforce simp add: inj_on_def) +lemma inj_on_image_eq_iff: "inj_on f C \ A \ C \ B \ C \ f ` A = f ` B \ A = B" + by (fastforce simp add: inj_on_def) lemma inj_on_Un_image_eq_iff: "inj_on f (A \ B) \ f ` A = f ` B \ A = B" -by(erule inj_on_image_eq_iff) simp_all + by (erule inj_on_image_eq_iff) simp_all -lemma inj_on_image_Int: - "[| inj_on f C; A<=C; B<=C |] ==> f`(A Int B) = f`A Int f`B" - by (simp add: inj_on_def, blast) +lemma inj_on_image_Int: "inj_on f C \ A \ C \ B \ C \ f ` (A \ B) = f ` A \ f ` B" + unfolding inj_on_def by blast + +lemma inj_on_image_set_diff: "inj_on f C \ A - B \ C \ B \ C \ f ` (A - B) = f ` A - f ` B" + unfolding inj_on_def by blast -lemma inj_on_image_set_diff: - "[| inj_on f C; A-B \ C; B \ C |] ==> f`(A-B) = f`A - f`B" - by (simp add: inj_on_def, blast) +lemma image_Int: "inj f \ f ` (A \ B) = f ` A \ f ` B" + unfolding inj_on_def by blast -lemma image_Int: "inj f ==> f`(A Int B) = f`A Int f`B" - by (simp add: inj_on_def, blast) +lemma image_set_diff: "inj f \ f ` (A - B) = f ` A - f ` B" + unfolding inj_on_def by blast -lemma image_set_diff: "inj f ==> f`(A-B) = f`A - f`B" -by (simp add: inj_on_def, blast) - -lemma inj_on_image_mem_iff: "\inj_on f B; a \ B; A \ B\ \ f a \ f`A \ a \ A" +lemma inj_on_image_mem_iff: "inj_on f B \ a \ B \ A \ B \ f a \ f ` A \ a \ A" by (auto simp: inj_on_def) (*FIXME DELETE*) -lemma inj_on_image_mem_iff_alt: "inj_on f B \ A \ B \ f a \ f`A \ a \ B \ a \ A" +lemma inj_on_image_mem_iff_alt: "inj_on f B \ A \ B \ f a \ f ` A \ a \ B \ a \ A" by (blast dest: inj_onD) -lemma inj_image_mem_iff: "inj f \ f a \ f`A \ a \ A" +lemma inj_image_mem_iff: "inj f \ f a \ f ` A \ a \ A" by (blast dest: injD) -lemma inj_image_subset_iff: "inj f ==> (f`A <= f`B) = (A<=B)" +lemma inj_image_subset_iff: "inj f \ f ` A \ f ` B \ A \ B" by (blast dest: injD) -lemma inj_image_eq_iff: "inj f ==> (f`A = f`B) = (A = B)" +lemma inj_image_eq_iff: "inj f \ f ` A = f ` B \ A = B" by (blast dest: injD) -lemma surj_Compl_image_subset: "surj f ==> -(f`A) <= f`(-A)" -by auto - -lemma inj_image_Compl_subset: "inj f ==> f`(-A) <= -(f`A)" -by (auto simp add: inj_on_def) +lemma surj_Compl_image_subset: "surj f \ - (f ` A) \ f ` (- A)" + by auto -lemma bij_image_Compl_eq: "bij f ==> f`(-A) = -(f`A)" -apply (simp add: bij_def) -apply (rule equalityI) -apply (simp_all (no_asm_simp) add: inj_image_Compl_subset surj_Compl_image_subset) -done +lemma inj_image_Compl_subset: "inj f \ f ` (- A) \ - (f ` A)" + by (auto simp add: inj_on_def) + +lemma bij_image_Compl_eq: "bij f \ f ` (- A) = - (f ` A)" + by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) lemma inj_vimage_singleton: "inj f \ f -` {a} \ {THE x. f x = a}" - \ \The inverse image of a singleton under an injective function - is included in a singleton.\ + \ \The inverse image of a singleton under an injective function is included in a singleton.\ apply (auto simp add: inj_on_def) apply (blast intro: the_equality [symmetric]) done -lemma inj_on_vimage_singleton: - "inj_on f A \ f -` {a} \ A \ {THE x. x \ A \ f x = a}" +lemma inj_on_vimage_singleton: "inj_on f A \ f -` {a} \ A \ {THE x. x \ A \ f x = a}" by (auto simp add: inj_on_def intro: the_equality [symmetric]) lemma (in ordered_ab_group_add) inj_uminus[simp, intro]: "inj_on uminus A" @@ -548,84 +511,92 @@ by (auto intro!: inj_onI dest: strict_mono_eq) lemma bij_betw_byWitness: -assumes LEFT: "\a \ A. f'(f a) = a" and - RIGHT: "\a' \ A'. f(f' a') = a'" and - IM1: "f ` A \ A'" and IM2: "f' ` A' \ A" -shows "bij_betw f A A'" -using assms -proof(unfold bij_betw_def inj_on_def, safe) - fix a b assume *: "a \ A" "b \ A" and **: "f a = f b" - have "a = f'(f a) \ b = f'(f b)" using * LEFT by simp + assumes left: "\a \ A. f' (f a) = a" + and right: "\a' \ A'. f (f' a') = a'" + and "f ` A \ A'" + and img2: "f' ` A' \ A" + shows "bij_betw f A A'" + using assms +proof (unfold bij_betw_def inj_on_def, safe) + fix a b + assume *: "a \ A" "b \ A" and **: "f a = f b" + have "a = f' (f a) \ b = f'(f b)" using * left by simp with ** show "a = b" by simp next fix a' assume *: "a' \ A'" - hence "f' a' \ A" using IM2 by blast + hence "f' a' \ A" using img2 by blast moreover - have "a' = f(f' a')" using * RIGHT by simp + have "a' = f (f' a')" using * right by simp ultimately show "a' \ f ` A" by blast qed corollary notIn_Un_bij_betw: -assumes NIN: "b \ A" and NIN': "f b \ A'" and - BIJ: "bij_betw f A A'" -shows "bij_betw f (A \ {b}) (A' \ {f b})" -proof- + assumes "b \ A" + and "f b \ A'" + and "bij_betw f A A'" + shows "bij_betw f (A \ {b}) (A' \ {f b})" +proof - have "bij_betw f {b} {f b}" - unfolding bij_betw_def inj_on_def by simp + unfolding bij_betw_def inj_on_def by simp with assms show ?thesis - using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast + using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast qed lemma notIn_Un_bij_betw3: -assumes NIN: "b \ A" and NIN': "f b \ A'" -shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" + assumes "b \ A" + and "f b \ A'" + shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" proof assume "bij_betw f A A'" - thus "bij_betw f (A \ {b}) (A' \ {f b})" - using assms notIn_Un_bij_betw[of b A f A'] by blast + then show "bij_betw f (A \ {b}) (A' \ {f b})" + using assms notIn_Un_bij_betw [of b A f A'] by blast next assume *: "bij_betw f (A \ {b}) (A' \ {f b})" have "f ` A = A'" - proof(auto) - fix a assume **: "a \ A" - hence "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast + proof auto + fix a + assume **: "a \ A" + then have "f a \ A' \ {f b}" + using * unfolding bij_betw_def by blast moreover - {assume "f a = f b" - hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast - with NIN ** have False by blast - } + have False if "f a = f b" + proof - + have "a = b" using * ** that unfolding bij_betw_def inj_on_def by blast + with \b \ A\ ** show ?thesis by blast + qed ultimately show "f a \ A'" by blast next - fix a' assume **: "a' \ A'" - hence "a' \ f`(A \ {b})" - using * by (auto simp add: bij_betw_def) + fix a' + assume **: "a' \ A'" + then have "a' \ f ` (A \ {b})" + using * by (auto simp add: bij_betw_def) then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast moreover - {assume "a = b" with 1 ** NIN' have False by blast - } + have False if "a = b" using 1 ** \f b \ A'\ that by blast ultimately have "a \ A" by blast with 1 show "a' \ f ` A" by blast qed - thus "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast + then show "bij_betw f A A'" + using * bij_betw_subset[of f "A \ {b}" _ A] by blast qed -subsection\Function Updating\ +subsection \Function Updating\ -definition fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where - "fun_upd f a b == % x. if x=a then b else f x" +definition fun_upd :: "('a \ 'b) \ 'a \ 'b \ ('a \ 'b)" + where "fun_upd f a b \ \x. if x = a then b else f x" nonterminal updbinds and updbind syntax - "_updbind" :: "['a, 'a] => updbind" ("(2_ :=/ _)") - "" :: "updbind => updbinds" ("_") - "_updbinds":: "[updbind, updbinds] => updbinds" ("_,/ _") - "_Update" :: "['a, updbinds] => 'a" ("_/'((_)')" [1000, 0] 900) + "_updbind" :: "'a \ 'a \ updbind" ("(2_ :=/ _)") + "" :: "updbind \ updbinds" ("_") + "_updbinds":: "updbind \ updbinds \ updbinds" ("_,/ _") + "_Update" :: "'a \ updbinds \ 'a" ("_/'((_)')" [1000, 0] 900) translations - "_Update f (_updbinds b bs)" == "_Update (_Update f b) bs" - "f(x:=y)" == "CONST fun_upd f x y" + "_Update f (_updbinds b bs)" \ "_Update (_Update f b) bs" + "f(x:=y)" \ "CONST fun_upd f x y" (* Hint: to define the sum of two functions (or maps), use case_sum. A nice infix syntax could be defined by @@ -633,69 +604,69 @@ case_sum (infixr "'(+')"80) *) -lemma fun_upd_idem_iff: "(f(x:=y) = f) = (f x = y)" -apply (simp add: fun_upd_def, safe) -apply (erule subst) -apply (rule_tac [2] ext, auto) -done +lemma fun_upd_idem_iff: "f(x:=y) = f \ f x = y" + unfolding fun_upd_def + apply safe + apply (erule subst) + apply (rule_tac [2] ext) + apply auto + done -lemma fun_upd_idem: "f x = y ==> f(x:=y) = f" +lemma fun_upd_idem: "f x = y \ f(x := y) = f" by (simp only: fun_upd_idem_iff) lemma fun_upd_triv [iff]: "f(x := f x) = f" by (simp only: fun_upd_idem) -lemma fun_upd_apply [simp]: "(f(x:=y))z = (if z=x then y else f z)" -by (simp add: fun_upd_def) +lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" + by (simp add: fun_upd_def) -(* fun_upd_apply supersedes these two, but they are useful +(* fun_upd_apply supersedes these two, but they are useful if fun_upd_apply is intentionally removed from the simpset *) -lemma fun_upd_same: "(f(x:=y)) x = y" -by simp +lemma fun_upd_same: "(f(x := y)) x = y" + by simp -lemma fun_upd_other: "z~=x ==> (f(x:=y)) z = f z" -by simp - -lemma fun_upd_upd [simp]: "f(x:=y,x:=z) = f(x:=z)" -by (simp add: fun_eq_iff) +lemma fun_upd_other: "z \ x \ (f(x := y)) z = f z" + by simp -lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)" -by (rule ext, auto) +lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" + by (simp add: fun_eq_iff) -lemma inj_on_fun_updI: - "inj_on f A \ y \ f ` A \ inj_on (f(x := y)) A" +lemma fun_upd_twist: "a \ c \ (m(a := b))(c := d) = (m(c := d))(a := b)" + by (rule ext) auto + +lemma inj_on_fun_updI: "inj_on f A \ y \ f ` A \ inj_on (f(x := y)) A" by (fastforce simp: inj_on_def) -lemma fun_upd_image: - "f(x:=y) ` A = (if x \ A then insert y (f ` (A-{x})) else f ` A)" -by auto +lemma fun_upd_image: "f(x := y) ` A = (if x \ A then insert y (f ` (A - {x})) else f ` A)" + by auto lemma fun_upd_comp: "f \ (g(x := y)) = (f \ g)(x := f y)" by auto lemma fun_upd_eqD: "f(x := y) = g(x := z) \ y = z" -by(simp add: fun_eq_iff split: if_split_asm) + by (simp add: fun_eq_iff split: if_split_asm) + subsection \\override_on\\ -definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" where - "override_on f g A = (\a. if a \ A then g a else f a)" +definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" + where "override_on f g A = (\a. if a \ A then g a else f a)" lemma override_on_emptyset[simp]: "override_on f g {} = f" -by(simp add:override_on_def) + by (simp add:override_on_def) -lemma override_on_apply_notin[simp]: "a ~: A ==> (override_on f g A) a = f a" -by(simp add:override_on_def) +lemma override_on_apply_notin[simp]: "a \ A \ (override_on f g A) a = f a" + by (simp add:override_on_def) -lemma override_on_apply_in[simp]: "a : A ==> (override_on f g A) a = g a" -by(simp add:override_on_def) +lemma override_on_apply_in[simp]: "a \ A \ (override_on f g A) a = g a" + by (simp add:override_on_def) subsection \\swap\\ definition swap :: "'a \ 'a \ ('a \ 'b) \ ('a \ 'b)" -where - "swap a b f = f (a := f b, b:= f a)" + where "swap a b f = f (a := f b, b:= f a)" lemma swap_apply [simp]: "swap a b f a = f b" @@ -703,20 +674,16 @@ "c \ a \ c \ b \ swap a b f c = f c" by (simp_all add: swap_def) -lemma swap_self [simp]: - "swap a a f = f" +lemma swap_self [simp]: "swap a a f = f" by (simp add: swap_def) -lemma swap_commute: - "swap a b f = swap b a f" +lemma swap_commute: "swap a b f = swap b a f" by (simp add: fun_upd_def swap_def fun_eq_iff) -lemma swap_nilpotent [simp]: - "swap a b (swap a b f) = f" +lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" by (rule ext, simp add: fun_upd_def swap_def) -lemma swap_comp_involutory [simp]: - "swap a b \ swap a b = id" +lemma swap_comp_involutory [simp]: "swap a b \ swap a b = id" by (rule ext) simp lemma swap_triple: @@ -725,10 +692,11 @@ using assms by (simp add: fun_eq_iff swap_def) lemma comp_swap: "f \ swap a b g = swap a b (f \ g)" - by (rule ext, simp add: fun_upd_def swap_def) + by (rule ext) (simp add: fun_upd_def swap_def) lemma swap_image_eq [simp]: - assumes "a \ A" "b \ A" shows "swap a b f ` A = f ` A" + assumes "a \ A" "b \ A" + shows "swap a b f ` A = f ` A" proof - have subset: "\f. swap a b f ` A \ f ` A" using assms by (auto simp: image_iff swap_def) @@ -736,20 +704,21 @@ with subset[of f] show ?thesis by auto qed -lemma inj_on_imp_inj_on_swap: - "\inj_on f A; a \ A; b \ A\ \ inj_on (swap a b f) A" - by (simp add: inj_on_def swap_def, blast) +lemma inj_on_imp_inj_on_swap: "inj_on f A \ a \ A \ b \ A \ inj_on (swap a b f) A" + by (auto simp add: inj_on_def swap_def) lemma inj_on_swap_iff [simp]: - assumes A: "a \ A" "b \ A" shows "inj_on (swap a b f) A \ inj_on f A" + assumes A: "a \ A" "b \ A" + shows "inj_on (swap a b f) A \ inj_on f A" proof assume "inj_on (swap a b f) A" with A have "inj_on (swap a b (swap a b f)) A" by (iprover intro: inj_on_imp_inj_on_swap) - thus "inj_on f A" by simp + then show "inj_on f A" by simp next assume "inj_on f A" - with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) + with A show "inj_on (swap a b f) A" + by (iprover intro: inj_on_imp_inj_on_swap) qed lemma surj_imp_surj_swap: "surj f \ surj (swap a b f)" @@ -758,8 +727,7 @@ lemma surj_swap_iff [simp]: "surj (swap a b f) \ surj f" by simp -lemma bij_betw_swap_iff [simp]: - "\ x \ A; y \ A \ \ bij_betw (swap x y f) A B \ bij_betw f A B" +lemma bij_betw_swap_iff [simp]: "x \ A \ y \ A \ bij_betw (swap x y f) A B \ bij_betw f A B" by (auto simp: bij_betw_def) lemma bij_swap_iff [simp]: "bij (swap a b f) \ bij f" @@ -770,114 +738,107 @@ subsection \Inversion of injective functions\ -definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where - "the_inv_into A f == %x. THE y. y : A & f y = x" +definition the_inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" + where "the_inv_into A f \ \x. THE y. y \ A \ f y = x" + +lemma the_inv_into_f_f: "inj_on f A \ x \ A \ the_inv_into A f (f x) = x" + unfolding the_inv_into_def inj_on_def by blast -lemma the_inv_into_f_f: - "[| inj_on f A; x : A |] ==> the_inv_into A f (f x) = x" -apply (simp add: the_inv_into_def inj_on_def) -apply blast -done - -lemma f_the_inv_into_f: - "inj_on f A ==> y : f`A ==> f (the_inv_into A f y) = y" -apply (simp add: the_inv_into_def) -apply (rule the1I2) - apply(blast dest: inj_onD) -apply blast -done +lemma f_the_inv_into_f: "inj_on f A \ y \ f ` A \ f (the_inv_into A f y) = y" + apply (simp add: the_inv_into_def) + apply (rule the1I2) + apply(blast dest: inj_onD) + apply blast + done -lemma the_inv_into_into: - "[| inj_on f A; x : f ` A; A <= B |] ==> the_inv_into A f x : B" -apply (simp add: the_inv_into_def) -apply (rule the1I2) - apply(blast dest: inj_onD) -apply blast -done +lemma the_inv_into_into: "inj_on f A \ x \ f ` A \ A \ B \ the_inv_into A f x \ B" + apply (simp add: the_inv_into_def) + apply (rule the1I2) + apply(blast dest: inj_onD) + apply blast + done -lemma the_inv_into_onto[simp]: - "inj_on f A ==> the_inv_into A f ` (f ` A) = A" -by (fast intro:the_inv_into_into the_inv_into_f_f[symmetric]) +lemma the_inv_into_onto [simp]: "inj_on f A \ the_inv_into A f ` (f ` A) = A" + by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) -lemma the_inv_into_f_eq: - "[| inj_on f A; f x = y; x : A |] ==> the_inv_into A f y = x" +lemma the_inv_into_f_eq: "inj_on f A \ f x = y \ x \ A \ the_inv_into A f y = x" apply (erule subst) - apply (erule the_inv_into_f_f, assumption) + apply (erule the_inv_into_f_f) + apply assumption done lemma the_inv_into_comp: - "[| inj_on f (g ` A); inj_on g A; x : f ` g ` A |] ==> - the_inv_into A (f o g) x = (the_inv_into A g o the_inv_into (g ` A) f) x" -apply (rule the_inv_into_f_eq) - apply (fast intro: comp_inj_on) - apply (simp add: f_the_inv_into_f the_inv_into_into) -apply (simp add: the_inv_into_into) -done + "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ + the_inv_into A (f \ g) x = (the_inv_into A g \ the_inv_into (g ` A) f) x" + apply (rule the_inv_into_f_eq) + apply (fast intro: comp_inj_on) + apply (simp add: f_the_inv_into_f the_inv_into_into) + apply (simp add: the_inv_into_into) + done -lemma inj_on_the_inv_into: - "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" -by (auto intro: inj_onI simp: the_inv_into_f_f) +lemma inj_on_the_inv_into: "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" + by (auto intro: inj_onI simp: the_inv_into_f_f) -lemma bij_betw_the_inv_into: - "bij_betw f A B \ bij_betw (the_inv_into A f) B A" -by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) +lemma bij_betw_the_inv_into: "bij_betw f A B \ bij_betw (the_inv_into A f) B A" + by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) -abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" where - "the_inv f \ the_inv_into UNIV f" +abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" + where "the_inv f \ the_inv_into UNIV f" lemma the_inv_f_f: assumes "inj f" - shows "the_inv f (f x) = x" using assms UNIV_I - by (rule the_inv_into_f_f) + shows "the_inv f (f x) = x" + using assms UNIV_I by (rule the_inv_into_f_f) subsection \Cantor's Paradox\ -lemma Cantors_paradox: - "\(\f. f ` A = Pow A)" +lemma Cantors_paradox: "\ (\f. f ` A = Pow A)" proof clarify - fix f assume "f ` A = Pow A" hence *: "Pow A \ f ` A" by blast + fix f + assume "f ` A = Pow A" + then have *: "Pow A \ f ` A" by blast let ?X = "{a \ A. a \ f a}" have "?X \ Pow A" unfolding Pow_def by auto with * obtain x where "x \ A \ f x = ?X" by blast - thus False by best + then show False by best qed + subsection \Setup\ subsubsection \Proof tools\ -text \simplifies terms of the form - f(...,x:=y,...,x:=z,...) to f(...,x:=z,...)\ +text \Simplify terms of the form f(...,x:=y,...,x:=z,...) to f(...,x:=z,...)\ simproc_setup fun_upd2 ("f(v := w, x := y)") = \fn _ => -let - fun gen_fun_upd NONE T _ _ = NONE - | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y) - fun dest_fun_T1 (Type (_, T :: Ts)) = T - fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) = - let - fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) = - if v aconv x then SOME g else gen_fun_upd (find g) T v w - | find t = NONE - in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end + let + fun gen_fun_upd NONE T _ _ = NONE + | gen_fun_upd (SOME f) T x y = SOME (Const (@{const_name fun_upd}, T) $ f $ x $ y) + fun dest_fun_T1 (Type (_, T :: Ts)) = T + fun find_double (t as Const (@{const_name fun_upd},T) $ f $ x $ y) = + let + fun find (Const (@{const_name fun_upd},T) $ g $ v $ w) = + if v aconv x then SOME g else gen_fun_upd (find g) T v w + | find t = NONE + in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end - val ss = simpset_of @{context} + val ss = simpset_of @{context} - fun proc ctxt ct = - let - val t = Thm.term_of ct - in - case find_double t of - (T, NONE) => NONE - | (T, SOME rhs) => - SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) - (fn _ => - resolve_tac ctxt [eq_reflection] 1 THEN - resolve_tac ctxt @{thms ext} 1 THEN - simp_tac (put_simpset ss ctxt) 1)) - end -in proc end + fun proc ctxt ct = + let + val t = Thm.term_of ct + in + case find_double t of + (T, NONE) => NONE + | (T, SOME rhs) => + SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) + (fn _ => + resolve_tac ctxt [eq_reflection] 1 THEN + resolve_tac ctxt @{thms ext} 1 THEN + simp_tac (put_simpset ss ctxt) 1)) + end + in proc end \ @@ -891,6 +852,7 @@ functor vimage by (simp_all add: fun_eq_iff vimage_comp) + text \Legacy theorem names\ lemmas o_def = comp_def @@ -904,4 +866,3 @@ lemmas o_eq_id_dest = comp_eq_id_dest end - diff -r dff40165618c -r bc1f17d45e91 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Sun Jun 19 22:51:42 2016 +0200 +++ b/src/HOL/Lattices.thy Mon Jun 20 17:03:50 2016 +0200 @@ -21,24 +21,23 @@ begin lemma left_idem [simp]: "a \<^bold>* (a \<^bold>* b) = a \<^bold>* b" -by (simp add: assoc [symmetric]) + by (simp add: assoc [symmetric]) lemma right_idem [simp]: "(a \<^bold>* b) \<^bold>* b = a \<^bold>* b" -by (simp add: assoc) + by (simp add: assoc) end locale semilattice_neutr = semilattice + comm_monoid locale semilattice_order = semilattice + - fixes less_eq :: "'a \ 'a \ bool" (infix "\<^bold>\" 50) - and less :: "'a \ 'a \ bool" (infix "\<^bold><" 50) + fixes less_eq :: "'a \ 'a \ bool" (infix "\<^bold>\" 50) + and less :: "'a \ 'a \ bool" (infix "\<^bold><" 50) assumes order_iff: "a \<^bold>\ b \ a = a \<^bold>* b" and strict_order_iff: "a \<^bold>< b \ a = a \<^bold>* b \ a \ b" begin -lemma orderI: - "a = a \<^bold>* b \ a \<^bold>\ b" +lemma orderI: "a = a \<^bold>* b \ a \<^bold>\ b" by (simp add: order_iff) lemma orderE: @@ -49,7 +48,7 @@ sublocale ordering less_eq less proof fix a b - show "a \<^bold>< b \ a \<^bold>\ b \ a \ b" + show "a \<^bold>< b \ a \<^bold>\ b \ a \ b" for a b by (simp add: order_iff strict_order_iff) next fix a @@ -74,12 +73,10 @@ then show "a \<^bold>\ c" by (rule orderI) qed -lemma cobounded1 [simp]: - "a \<^bold>* b \<^bold>\ a" - by (simp add: order_iff commute) +lemma cobounded1 [simp]: "a \<^bold>* b \<^bold>\ a" + by (simp add: order_iff commute) -lemma cobounded2 [simp]: - "a \<^bold>* b \<^bold>\ b" +lemma cobounded2 [simp]: "a \<^bold>* b \<^bold>\ b" by (simp add: order_iff) lemma boundedI: @@ -95,8 +92,7 @@ obtains "a \<^bold>\ b" and "a \<^bold>\ c" using assms by (blast intro: trans cobounded1 cobounded2) -lemma bounded_iff [simp]: - "a \<^bold>\ b \<^bold>* c \ a \<^bold>\ b \ a \<^bold>\ c" +lemma bounded_iff [simp]: "a \<^bold>\ b \<^bold>* c \ a \<^bold>\ b \ a \<^bold>\ c" by (blast intro: boundedI elim: boundedE) lemma strict_boundedE: @@ -104,21 +100,17 @@ obtains "a \<^bold>< b" and "a \<^bold>< c" using assms by (auto simp add: commute strict_iff_order elim: orderE intro!: that)+ -lemma coboundedI1: - "a \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" +lemma coboundedI1: "a \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" by (rule trans) auto -lemma coboundedI2: - "b \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" +lemma coboundedI2: "b \<^bold>\ c \ a \<^bold>* b \<^bold>\ c" by (rule trans) auto -lemma strict_coboundedI1: - "a \<^bold>< c \ a \<^bold>* b \<^bold>< c" +lemma strict_coboundedI1: "a \<^bold>< c \ a \<^bold>* b \<^bold>< c" using irrefl by (auto intro: not_eq_order_implies_strict coboundedI1 strict_implies_order elim: strict_boundedE) -lemma strict_coboundedI2: - "b \<^bold>< c \ a \<^bold>* b \<^bold>< c" +lemma strict_coboundedI2: "b \<^bold>< c \ a \<^bold>* b \<^bold>< c" using strict_coboundedI1 [of b c a] by (simp add: commute) lemma mono: "a \<^bold>\ c \ b \<^bold>\ d \ a \<^bold>* b \<^bold>\ c \<^bold>* d" @@ -152,7 +144,7 @@ class inf = fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) -class sup = +class sup = fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) @@ -175,10 +167,9 @@ text \Dual lattice\ -lemma dual_semilattice: - "class.semilattice_inf sup greater_eq greater" -by (rule class.semilattice_inf.intro, rule dual_order) - (unfold_locales, simp_all add: sup_least) +lemma dual_semilattice: "class.semilattice_inf sup greater_eq greater" + by (rule class.semilattice_inf.intro, rule dual_order) + (unfold_locales, simp_all add: sup_least) end @@ -190,12 +181,10 @@ context semilattice_inf begin -lemma le_infI1: - "a \ x \ a \ b \ x" +lemma le_infI1: "a \ x \ a \ b \ x" by (rule order_trans) auto -lemma le_infI2: - "b \ x \ a \ b \ x" +lemma le_infI2: "b \ x \ a \ b \ x" by (rule order_trans) auto lemma le_infI: "x \ a \ x \ b \ x \ a \ b" @@ -204,20 +193,16 @@ lemma le_infE: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" by (blast intro: order_trans inf_le1 inf_le2) -lemma le_inf_iff: - "x \ y \ z \ x \ y \ x \ z" +lemma le_inf_iff: "x \ y \ z \ x \ y \ x \ z" by (blast intro: le_infI elim: le_infE) -lemma le_iff_inf: - "x \ y \ x \ y = x" +lemma le_iff_inf: "x \ y \ x \ y = x" by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1] simp add: le_inf_iff) lemma inf_mono: "a \ c \ b \ d \ a \ b \ c \ d" by (fast intro: inf_greatest le_infI1 le_infI2) -lemma mono_inf: - fixes f :: "'a \ 'b::semilattice_inf" - shows "mono f \ f (A \ B) \ f A \ f B" +lemma mono_inf: "mono f \ f (A \ B) \ f A \ f B" for f :: "'a \ 'b::semilattice_inf" by (auto simp add: mono_def intro: Lattices.inf_greatest) end @@ -225,36 +210,28 @@ context semilattice_sup begin -lemma le_supI1: - "x \ a \ x \ a \ b" +lemma le_supI1: "x \ a \ x \ a \ b" + by (rule order_trans) auto + +lemma le_supI2: "x \ b \ x \ a \ b" by (rule order_trans) auto -lemma le_supI2: - "x \ b \ x \ a \ b" - by (rule order_trans) auto - -lemma le_supI: - "a \ x \ b \ x \ a \ b \ x" +lemma le_supI: "a \ x \ b \ x \ a \ b \ x" by (fact sup_least) (* FIXME: duplicate lemma *) -lemma le_supE: - "a \ b \ x \ (a \ x \ b \ x \ P) \ P" +lemma le_supE: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" by (blast intro: order_trans sup_ge1 sup_ge2) -lemma le_sup_iff: - "x \ y \ z \ x \ z \ y \ z" +lemma le_sup_iff: "x \ y \ z \ x \ z \ y \ z" by (blast intro: le_supI elim: le_supE) -lemma le_iff_sup: - "x \ y \ x \ y = y" +lemma le_iff_sup: "x \ y \ x \ y = y" by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1] simp add: le_sup_iff) lemma sup_mono: "a \ c \ b \ d \ a \ b \ c \ d" by (fast intro: sup_least le_supI1 le_supI2) -lemma mono_sup: - fixes f :: "'a \ 'b::semilattice_sup" - shows "mono f \ f A \ f B \ f (A \ B)" +lemma mono_sup: "mono f \ f A \ f B \ f (A \ B)" for f :: "'a \ 'b::semilattice_sup" by (auto simp add: mono_def intro: Lattices.sup_least) end @@ -302,7 +279,7 @@ lemma inf_absorb2: "y \ x \ x \ y = y" by (rule antisym) auto - + lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem end @@ -352,8 +329,7 @@ context lattice begin -lemma dual_lattice: - "class.lattice sup (op \) (op >) inf" +lemma dual_lattice: "class.lattice sup (op \) (op >) inf" by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order) (unfold_locales, auto) @@ -375,47 +351,48 @@ lemma distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) -text\If you have one of them, you have them all.\ +text \If you have one of them, you have them all.\ lemma distrib_imp1: -assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" -shows "x \ (y \ z) = (x \ y) \ (x \ z)" + assumes distrib: "\x y z. x \ (y \ z) = (x \ y) \ (x \ z)" + shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- - have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp + have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" + by simp also have "\ = x \ (z \ (x \ y))" - by (simp add: D inf_commute sup_assoc del: sup_inf_absorb) + by (simp add: distrib inf_commute sup_assoc del: sup_inf_absorb) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" - by(simp add: inf_commute) - also have "\ = (x \ y) \ (x \ z)" by(simp add:D) + by (simp add: inf_commute) + also have "\ = (x \ y) \ (x \ z)" by(simp add:distrib) finally show ?thesis . qed lemma distrib_imp2: -assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" -shows "x \ (y \ z) = (x \ y) \ (x \ z)" + assumes distrib: "\x y z. x \ (y \ z) = (x \ y) \ (x \ z)" + shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- - have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp + have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" + by simp also have "\ = x \ (z \ (x \ y))" - by (simp add: D sup_commute inf_assoc del: inf_sup_absorb) + by (simp add: distrib sup_commute inf_assoc del: inf_sup_absorb) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" - by(simp add: sup_commute) - also have "\ = (x \ y) \ (x \ z)" by(simp add:D) + by (simp add: sup_commute) + also have "\ = (x \ y) \ (x \ z)" by (simp add:distrib) finally show ?thesis . qed end + subsubsection \Strict order\ context semilattice_inf begin -lemma less_infI1: - "a \ x \ a \ b \ x" +lemma less_infI1: "a \ x \ a \ b \ x" by (auto simp add: less_le inf_absorb1 intro: le_infI1) -lemma less_infI2: - "b \ x \ a \ b \ x" +lemma less_infI2: "b \ x \ a \ b \ x" by (auto simp add: less_le inf_absorb2 intro: le_infI2) end @@ -423,13 +400,11 @@ context semilattice_sup begin -lemma less_supI1: - "x \ a \ x \ a \ b" +lemma less_supI1: "x \ a \ x \ a \ b" using dual_semilattice by (rule semilattice_inf.less_infI1) -lemma less_supI2: - "x \ b \ x \ a \ b" +lemma less_supI2: "x \ b \ x \ a \ b" using dual_semilattice by (rule semilattice_inf.less_infI2) @@ -444,31 +419,24 @@ context distrib_lattice begin -lemma sup_inf_distrib2: - "(y \ z) \ x = (y \ x) \ (z \ x)" +lemma sup_inf_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by (simp add: sup_commute sup_inf_distrib1) -lemma inf_sup_distrib1: - "x \ (y \ z) = (x \ y) \ (x \ z)" +lemma inf_sup_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" by (rule distrib_imp2 [OF sup_inf_distrib1]) -lemma inf_sup_distrib2: - "(y \ z) \ x = (y \ x) \ (z \ x)" +lemma inf_sup_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by (simp add: inf_commute inf_sup_distrib1) -lemma dual_distrib_lattice: - "class.distrib_lattice sup (op \) (op >) inf" +lemma dual_distrib_lattice: "class.distrib_lattice sup (op \) (op >) inf" by (rule class.distrib_lattice.intro, rule dual_lattice) (unfold_locales, fact inf_sup_distrib1) -lemmas sup_inf_distrib = - sup_inf_distrib1 sup_inf_distrib2 +lemmas sup_inf_distrib = sup_inf_distrib1 sup_inf_distrib2 -lemmas inf_sup_distrib = - inf_sup_distrib1 inf_sup_distrib2 +lemmas inf_sup_distrib = inf_sup_distrib1 inf_sup_distrib2 -lemmas distrib = - sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 +lemmas distrib = sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 end @@ -481,8 +449,7 @@ sublocale inf_top: semilattice_neutr inf top + inf_top: semilattice_neutr_order inf top less_eq less proof - fix x - show "x \ \ = x" + show "x \ \ = x" for x by (rule inf_absorb1) simp qed @@ -494,8 +461,7 @@ sublocale sup_bot: semilattice_neutr sup bot + sup_bot: semilattice_neutr_order sup bot greater_eq greater proof - fix x - show "x \ \ = x" + show "x \ \ = x" for x by (rule sup_absorb1) simp qed @@ -506,28 +472,22 @@ subclass bounded_semilattice_sup_bot .. -lemma inf_bot_left [simp]: - "\ \ x = \" +lemma inf_bot_left [simp]: "\ \ x = \" by (rule inf_absorb1) simp -lemma inf_bot_right [simp]: - "x \ \ = \" +lemma inf_bot_right [simp]: "x \ \ = \" by (rule inf_absorb2) simp -lemma sup_bot_left: - "\ \ x = x" +lemma sup_bot_left: "\ \ x = x" by (fact sup_bot.left_neutral) -lemma sup_bot_right: - "x \ \ = x" +lemma sup_bot_right: "x \ \ = x" by (fact sup_bot.right_neutral) -lemma sup_eq_bot_iff [simp]: - "x \ y = \ \ x = \ \ y = \" +lemma sup_eq_bot_iff [simp]: "x \ y = \ \ x = \ \ y = \" by (simp add: eq_iff) -lemma bot_eq_sup_iff [simp]: - "\ = x \ y \ x = \ \ y = \" +lemma bot_eq_sup_iff [simp]: "\ = x \ y \ x = \ \ y = \" by (simp add: eq_iff) end @@ -537,24 +497,19 @@ subclass bounded_semilattice_inf_top .. -lemma sup_top_left [simp]: - "\ \ x = \" +lemma sup_top_left [simp]: "\ \ x = \" by (rule sup_absorb1) simp -lemma sup_top_right [simp]: - "x \ \ = \" +lemma sup_top_right [simp]: "x \ \ = \" by (rule sup_absorb2) simp -lemma inf_top_left: - "\ \ x = x" +lemma inf_top_left: "\ \ x = x" by (fact inf_top.left_neutral) -lemma inf_top_right: - "x \ \ = x" +lemma inf_top_right: "x \ \ = x" by (fact inf_top.right_neutral) -lemma inf_eq_top_iff [simp]: - "x \ y = \ \ x = \ \ y = \" +lemma inf_eq_top_iff [simp]: "x \ y = \ \ x = \ \ y = \" by (simp add: eq_iff) end @@ -565,8 +520,7 @@ subclass bounded_lattice_bot .. subclass bounded_lattice_top .. -lemma dual_bounded_lattice: - "class.bounded_lattice sup greater_eq greater inf \ \" +lemma dual_bounded_lattice: "class.bounded_lattice sup greater_eq greater inf \ \" by unfold_locales (auto simp add: less_le_not_le) end @@ -582,12 +536,10 @@ by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) -lemma compl_inf_bot [simp]: - "- x \ x = \" +lemma compl_inf_bot [simp]: "- x \ x = \" by (simp add: inf_commute inf_compl_bot) -lemma compl_sup_top [simp]: - "- x \ x = \" +lemma compl_sup_top [simp]: "- x \ x = \" by (simp add: sup_commute sup_compl_top) lemma compl_unique: @@ -606,12 +558,10 @@ then show "- x = y" by simp qed -lemma double_compl [simp]: - "- (- x) = x" +lemma double_compl [simp]: "- (- x) = x" using compl_inf_bot compl_sup_top by (rule compl_unique) -lemma compl_eq_compl_iff [simp]: - "- x = - y \ x = y" +lemma compl_eq_compl_iff [simp]: "- x = - y \ x = y" proof assume "- x = - y" then have "- (- x) = - (- y)" by (rule arg_cong) @@ -621,22 +571,19 @@ then show "- x = - y" by simp qed -lemma compl_bot_eq [simp]: - "- \ = \" +lemma compl_bot_eq [simp]: "- \ = \" proof - from sup_compl_top have "\ \ - \ = \" . then show ?thesis by simp qed -lemma compl_top_eq [simp]: - "- \ = \" +lemma compl_top_eq [simp]: "- \ = \" proof - from inf_compl_bot have "\ \ - \ = \" . then show ?thesis by simp qed -lemma compl_inf [simp]: - "- (x \ y) = - x \ - y" +lemma compl_inf [simp]: "- (x \ y) = - x \ - y" proof (rule compl_unique) have "(x \ y) \ (- x \ - y) = (y \ (x \ - x)) \ (x \ (y \ - y))" by (simp only: inf_sup_distrib inf_aci) @@ -649,86 +596,87 @@ by (simp add: sup_compl_top) qed -lemma compl_sup [simp]: - "- (x \ y) = - x \ - y" +lemma compl_sup [simp]: "- (x \ y) = - x \ - y" using dual_boolean_algebra by (rule boolean_algebra.compl_inf) lemma compl_mono: - "x \ y \ - y \ - x" + assumes "x \ y" + shows "- y \ - x" proof - - assume "x \ y" - then have "x \ y = y" by (simp only: le_iff_sup) + from assms have "x \ y = y" by (simp only: le_iff_sup) then have "- (x \ y) = - y" by simp then have "- x \ - y = - y" by simp then have "- y \ - x = - y" by (simp only: inf_commute) - then show "- y \ - x" by (simp only: le_iff_inf) + then show ?thesis by (simp only: le_iff_inf) qed -lemma compl_le_compl_iff [simp]: - "- x \ - y \ y \ x" +lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" by (auto dest: compl_mono) lemma compl_le_swap1: - assumes "y \ - x" shows "x \ -y" + assumes "y \ - x" + shows "x \ -y" proof - from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed lemma compl_le_swap2: - assumes "- y \ x" shows "- x \ y" + assumes "- y \ x" + shows "- x \ y" proof - from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed -lemma compl_less_compl_iff: (* TODO: declare [simp] ? *) - "- x \ - y \ y \ x" +lemma compl_less_compl_iff: "- x \ - y \ y \ x" (* TODO: declare [simp] ? *) by (auto simp add: less_le) lemma compl_less_swap1: - assumes "y \ - x" shows "x \ - y" + assumes "y \ - x" + shows "x \ - y" proof - from assms have "- (- x) \ - y" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma compl_less_swap2: - assumes "- y \ x" shows "- x \ y" + assumes "- y \ x" + shows "- x \ y" proof - from assms have "- x \ - (- y)" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma sup_cancel_left1: "sup (sup x a) (sup (- x) b) = top" -by(simp add: inf_sup_aci sup_compl_top) + by (simp add: inf_sup_aci sup_compl_top) lemma sup_cancel_left2: "sup (sup (- x) a) (sup x b) = top" -by(simp add: inf_sup_aci sup_compl_top) + by (simp add: inf_sup_aci sup_compl_top) lemma inf_cancel_left1: "inf (inf x a) (inf (- x) b) = bot" -by(simp add: inf_sup_aci inf_compl_bot) + by (simp add: inf_sup_aci inf_compl_bot) lemma inf_cancel_left2: "inf (inf (- x) a) (inf x b) = bot" -by(simp add: inf_sup_aci inf_compl_bot) + by (simp add: inf_sup_aci inf_compl_bot) -declare inf_compl_bot [simp] sup_compl_top [simp] +declare inf_compl_bot [simp] and sup_compl_top [simp] lemma sup_compl_top_left1 [simp]: "sup (- x) (sup x y) = top" -by(simp add: sup_assoc[symmetric]) + by (simp add: sup_assoc[symmetric]) lemma sup_compl_top_left2 [simp]: "sup x (sup (- x) y) = top" -using sup_compl_top_left1[of "- x" y] by simp + using sup_compl_top_left1[of "- x" y] by simp lemma inf_compl_bot_left1 [simp]: "inf (- x) (inf x y) = bot" -by(simp add: inf_assoc[symmetric]) + by (simp add: inf_assoc[symmetric]) lemma inf_compl_bot_left2 [simp]: "inf x (inf (- x) y) = bot" -using inf_compl_bot_left1[of "- x" y] by simp + using inf_compl_bot_left1[of "- x" y] by simp lemma inf_compl_bot_right [simp]: "inf x (inf y (- x)) = bot" -by(subst inf_left_commute) simp + by (subst inf_left_commute) simp end @@ -740,6 +688,7 @@ simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") = \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\ + subsection \\min/max\ as special case of lattice\ context linorder @@ -749,64 +698,48 @@ + max: semilattice_order max greater_eq greater by standard (auto simp add: min_def max_def) -lemma min_le_iff_disj: - "min x y \ z \ x \ z \ y \ z" +lemma min_le_iff_disj: "min x y \ z \ x \ z \ y \ z" unfolding min_def using linear by (auto intro: order_trans) -lemma le_max_iff_disj: - "z \ max x y \ z \ x \ z \ y" +lemma le_max_iff_disj: "z \ max x y \ z \ x \ z \ y" unfolding max_def using linear by (auto intro: order_trans) -lemma min_less_iff_disj: - "min x y < z \ x < z \ y < z" +lemma min_less_iff_disj: "min x y < z \ x < z \ y < z" unfolding min_def le_less using less_linear by (auto intro: less_trans) -lemma less_max_iff_disj: - "z < max x y \ z < x \ z < y" +lemma less_max_iff_disj: "z < max x y \ z < x \ z < y" unfolding max_def le_less using less_linear by (auto intro: less_trans) -lemma min_less_iff_conj [simp]: - "z < min x y \ z < x \ z < y" +lemma min_less_iff_conj [simp]: "z < min x y \ z < x \ z < y" unfolding min_def le_less using less_linear by (auto intro: less_trans) -lemma max_less_iff_conj [simp]: - "max x y < z \ x < z \ y < z" +lemma max_less_iff_conj [simp]: "max x y < z \ x < z \ y < z" unfolding max_def le_less using less_linear by (auto intro: less_trans) -lemma min_max_distrib1: - "min (max b c) a = max (min b a) (min c a)" +lemma min_max_distrib1: "min (max b c) a = max (min b a) (min c a)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) -lemma min_max_distrib2: - "min a (max b c) = max (min a b) (min a c)" +lemma min_max_distrib2: "min a (max b c) = max (min a b) (min a c)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) -lemma max_min_distrib1: - "max (min b c) a = min (max b a) (max c a)" +lemma max_min_distrib1: "max (min b c) a = min (max b a) (max c a)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) -lemma max_min_distrib2: - "max a (min b c) = min (max a b) (max a c)" +lemma max_min_distrib2: "max a (min b c) = min (max a b) (max a c)" by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) lemmas min_max_distribs = min_max_distrib1 min_max_distrib2 max_min_distrib1 max_min_distrib2 -lemma split_min [no_atp]: - "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" +lemma split_min [no_atp]: "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" by (simp add: min_def) -lemma split_max [no_atp]: - "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" +lemma split_max [no_atp]: "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" by (simp add: max_def) -lemma min_of_mono: - fixes f :: "'a \ 'b::linorder" - shows "mono f \ min (f m) (f n) = f (min m n)" +lemma min_of_mono: "mono f \ min (f m) (f n) = f (min m n)" for f :: "'a \ 'b::linorder" by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) -lemma max_of_mono: - fixes f :: "'a \ 'b::linorder" - shows "mono f \ max (f m) (f n) = f (max m n)" +lemma max_of_mono: "mono f \ max (f m) (f n) = f (max m n)" for f :: "'a \ 'b::linorder" by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) end @@ -821,27 +754,33 @@ subsection \Uniqueness of inf and sup\ lemma (in semilattice_inf) inf_unique: - fixes f (infixl "\" 70) - assumes le1: "\x y. x \ y \ x" and le2: "\x y. x \ y \ y" - and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" + fixes f (infixl "\" 70) + assumes le1: "\x y. x \ y \ x" + and le2: "\x y. x \ y \ y" + and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" shows "x \ y = x \ y" proof (rule antisym) - show "x \ y \ x \ y" by (rule le_infI) (rule le1, rule le2) -next - have leI: "\x y z. x \ y \ x \ z \ x \ y \ z" by (blast intro: greatest) - show "x \ y \ x \ y" by (rule leI) simp_all + show "x \ y \ x \ y" + by (rule le_infI) (rule le1, rule le2) + have leI: "\x y z. x \ y \ x \ z \ x \ y \ z" + by (blast intro: greatest) + show "x \ y \ x \ y" + by (rule leI) simp_all qed lemma (in semilattice_sup) sup_unique: - fixes f (infixl "\" 70) - assumes ge1 [simp]: "\x y. x \ x \ y" and ge2: "\x y. y \ x \ y" - and least: "\x y z. y \ x \ z \ x \ y \ z \ x" + fixes f (infixl "\" 70) + assumes ge1 [simp]: "\x y. x \ x \ y" + and ge2: "\x y. y \ x \ y" + and least: "\x y z. y \ x \ z \ x \ y \ z \ x" shows "x \ y = x \ y" proof (rule antisym) - show "x \ y \ x \ y" by (rule le_supI) (rule ge1, rule ge2) -next - have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" by (blast intro: least) - show "x \ y \ x \ y" by (rule leI) simp_all + show "x \ y \ x \ y" + by (rule le_supI) (rule ge1, rule ge2) + have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" + by (blast intro: least) + show "x \ y \ x \ y" + by (rule leI) simp_all qed @@ -850,33 +789,25 @@ instantiation bool :: boolean_algebra begin -definition - bool_Compl_def [simp]: "uminus = Not" +definition bool_Compl_def [simp]: "uminus = Not" -definition - bool_diff_def [simp]: "A - B \ A \ \ B" +definition bool_diff_def [simp]: "A - B \ A \ \ B" -definition - [simp]: "P \ Q \ P \ Q" +definition [simp]: "P \ Q \ P \ Q" -definition - [simp]: "P \ Q \ P \ Q" +definition [simp]: "P \ Q \ P \ Q" -instance proof -qed auto +instance by standard auto end -lemma sup_boolI1: - "P \ P \ Q" +lemma sup_boolI1: "P \ P \ Q" by simp -lemma sup_boolI2: - "Q \ P \ Q" +lemma sup_boolI2: "Q \ P \ Q" by simp -lemma sup_boolE: - "P \ Q \ (P \ R) \ (Q \ R) \ R" +lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" by auto @@ -885,48 +816,40 @@ instantiation "fun" :: (type, semilattice_sup) semilattice_sup begin -definition - "f \ g = (\x. f x \ g x)" +definition "f \ g = (\x. f x \ g x)" -lemma sup_apply [simp, code]: - "(f \ g) x = f x \ g x" +lemma sup_apply [simp, code]: "(f \ g) x = f x \ g x" by (simp add: sup_fun_def) -instance proof -qed (simp_all add: le_fun_def) +instance by standard (simp_all add: le_fun_def) end instantiation "fun" :: (type, semilattice_inf) semilattice_inf begin -definition - "f \ g = (\x. f x \ g x)" +definition "f \ g = (\x. f x \ g x)" -lemma inf_apply [simp, code]: - "(f \ g) x = f x \ g x" +lemma inf_apply [simp, code]: "(f \ g) x = f x \ g x" by (simp add: inf_fun_def) -instance proof -qed (simp_all add: le_fun_def) +instance by standard (simp_all add: le_fun_def) end instance "fun" :: (type, lattice) lattice .. -instance "fun" :: (type, distrib_lattice) distrib_lattice proof -qed (rule ext, simp add: sup_inf_distrib1) +instance "fun" :: (type, distrib_lattice) distrib_lattice + by standard (rule ext, simp add: sup_inf_distrib1) instance "fun" :: (type, bounded_lattice) bounded_lattice .. instantiation "fun" :: (type, uminus) uminus begin -definition - fun_Compl_def: "- A = (\x. - A x)" +definition fun_Compl_def: "- A = (\x. - A x)" -lemma uminus_apply [simp, code]: - "(- A) x = - (A x)" +lemma uminus_apply [simp, code]: "(- A) x = - (A x)" by (simp add: fun_Compl_def) instance .. @@ -936,19 +859,17 @@ instantiation "fun" :: (type, minus) minus begin -definition - fun_diff_def: "A - B = (\x. A x - B x)" +definition fun_diff_def: "A - B = (\x. A x - B x)" -lemma minus_apply [simp, code]: - "(A - B) x = A x - B x" +lemma minus_apply [simp, code]: "(A - B) x = A x - B x" by (simp add: fun_diff_def) instance .. end -instance "fun" :: (type, boolean_algebra) boolean_algebra proof -qed (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ +instance "fun" :: (type, boolean_algebra) boolean_algebra + by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ subsection \Lattice on unary and binary predicates\ @@ -995,10 +916,7 @@ lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" by (simp add: sup_fun_def) iprover -text \ - \medskip Classical introduction rule: no commitment to \A\ vs - \B\. -\ +text \ \<^medskip> Classical introduction rule: no commitment to \A\ vs \B\.\ lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" by (auto simp add: sup_fun_def) @@ -1012,4 +930,3 @@ less (infix "\" 50) end -