diff --git a/src/HOL/Conditionally_Complete_Lattices.thy b/src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy +++ b/src/HOL/Conditionally_Complete_Lattices.thy @@ -1,738 +1,739 @@ (* Title: HOL/Conditionally_Complete_Lattices.thy Author: Amine Chaieb and L C Paulson, University of Cambridge Author: Johannes Hölzl, TU München Author: Luke S. Serafin, Carnegie Mellon University *) section \Conditionally-complete Lattices\ theory Conditionally_Complete_Lattices imports Finite_Set Lattices_Big Set_Interval begin context preorder begin definition "bdd_above A \ (\M. \x \ A. x \ M)" definition "bdd_below A \ (\m. \x \ A. m \ x)" lemma bdd_aboveI[intro]: "(\x. x \ A \ x \ M) \ bdd_above A" by (auto simp: bdd_above_def) lemma bdd_belowI[intro]: "(\x. x \ A \ m \ x) \ bdd_below A" by (auto simp: bdd_below_def) lemma bdd_aboveI2: "(\x. x \ A \ f x \ M) \ bdd_above (f`A)" by force lemma bdd_belowI2: "(\x. x \ A \ m \ f x) \ bdd_below (f`A)" by force lemma bdd_above_empty [simp, intro]: "bdd_above {}" unfolding bdd_above_def by auto lemma bdd_below_empty [simp, intro]: "bdd_below {}" unfolding bdd_below_def by auto lemma bdd_above_mono: "bdd_above B \ A \ B \ bdd_above A" by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD) lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" by (metis bdd_below_def order_class.le_neq_trans psubsetD) lemma bdd_above_Int1 [simp]: "bdd_above A \ bdd_above (A \ B)" using bdd_above_mono by auto lemma bdd_above_Int2 [simp]: "bdd_above B \ bdd_above (A \ B)" using bdd_above_mono by auto lemma bdd_below_Int1 [simp]: "bdd_below A \ bdd_below (A \ B)" using bdd_below_mono by auto lemma bdd_below_Int2 [simp]: "bdd_below B \ bdd_below (A \ B)" using bdd_below_mono by auto lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) end lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A" by (rule bdd_aboveI[of _ top]) simp lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A" by (rule bdd_belowI[of _ bot]) simp lemma bdd_above_image_mono: "mono f \ bdd_above A \ bdd_above (f`A)" by (auto simp: bdd_above_def mono_def) lemma bdd_below_image_mono: "mono f \ bdd_below A \ bdd_below (f`A)" by (auto simp: bdd_below_def mono_def) lemma bdd_above_image_antimono: "antimono f \ bdd_below A \ bdd_above (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma bdd_below_image_antimono: "antimono f \ bdd_above A \ bdd_below (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma fixes X :: "'a::ordered_ab_group_add set" shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \ bdd_below X" and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \ bdd_above X" using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"] using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"] by (auto simp: antimono_def image_image) context lattice begin lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" by (auto simp: bdd_below_def intro: le_infI2 inf_le1) lemma bdd_finite [simp]: assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" using assms by (induct rule: finite_induct, auto) lemma bdd_above_Un [simp]: "bdd_above (A \ B) = (bdd_above A \ bdd_above B)" proof assume "bdd_above (A \ B)" thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto next assume "bdd_above A \ bdd_above B" then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) thus "bdd_above (A \ B)" unfolding bdd_above_def .. qed lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" proof assume "bdd_below (A \ B)" thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto next assume "bdd_below A \ bdd_below B" then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) thus "bdd_below (A \ B)" unfolding bdd_below_def .. qed lemma bdd_above_image_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" by (auto simp: bdd_above_def intro: le_supI1 le_supI2) lemma bdd_below_image_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" by (auto simp: bdd_below_def intro: le_infI1 le_infI2) lemma bdd_below_UN[simp]: "finite I \ bdd_below (\i\I. A i) = (\i \ I. bdd_below (A i))" by (induction I rule: finite.induct) auto lemma bdd_above_UN[simp]: "finite I \ bdd_above (\i\I. A i) = (\i \ I. bdd_above (A i))" by (induction I rule: finite.induct) auto end text \ To avoid name classes with the \<^class>\complete_lattice\-class we prefix \<^const>\Sup\ and \<^const>\Inf\ in theorem names with c. \ class conditionally_complete_lattice = lattice + Sup + Inf + assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" by (metis cSup_upper order_trans) lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" by (metis cInf_lower order_trans) lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" by (metis cSup_least cSup_upper2) lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" by (metis cInf_greatest cInf_lower2) lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" by (metis cSup_least cSup_upper subsetD) lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" by (metis cInf_greatest cInf_lower subsetD) lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" by (metis order_trans cSup_upper cSup_least) lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" by (metis order_trans cInf_lower cInf_greatest) lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ a" assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper) lemma cInf_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ a \ x" assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower) lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) lemma cSup_insert: "X \ {} \ bdd_above X \ Sup (insert a X) = sup a (Sup X)" by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) lemma cSup_singleton [simp]: "Sup {x} = x" by (intro cSup_eq_maximum) auto lemma cInf_singleton [simp]: "Inf {x} = x" by (intro cInf_eq_minimum) auto lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" using cSup_insert[of X] by simp lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" using cInf_insert[of X] by simp lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2) qed simp lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2) qed simp lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) lemma cSup_atMost[simp]: "Sup {..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" by (auto intro!: cSup_eq_maximum) lemma cInf_atLeast[simp]: "Inf {x..} = x" by (auto intro!: cInf_eq_minimum) lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" by (auto intro!: cInf_eq_minimum) lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ \(f ` A) \ f x" using cInf_lower [of _ "f ` A"] by simp lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ \(f ` A)" using cInf_greatest [of "f ` A"] by auto lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ \(f ` A)" using cSup_upper [of _ "f ` A"] by simp lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ \(f ` A) \ M" using cSup_least [of "f ` A"] by auto lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ \(f ` A) \ u" by (auto intro: cINF_lower order_trans) lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ \(f ` A)" by (auto intro: cSUP_upper order_trans) lemma cSUP_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cSUP_least) (auto intro: cSUP_upper) lemma cINF_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cINF_greatest) (auto intro: cINF_lower) lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ \(f ` A) \ (\x\A. u \ f x)" by (metis cINF_greatest cINF_lower order_trans) lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ \(f ` A) \ u \ (\x\A. f x \ u)" by (metis cSUP_least cSUP_upper order_trans) lemma less_cINF_D: "bdd_below (f`A) \ y < (\i\A. f i) \ i \ A \ y < f i" by (metis cINF_lower less_le_trans) lemma cSUP_lessD: "bdd_above (f`A) \ (\i\A. f i) < y \ i \ A \ f i < y" by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ \(f ` insert a A) = inf (f a) (\(f ` A))" by (simp add: cInf_insert cong del: INF_cong) lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ \(f ` insert a A) = sup (f a) (\(f ` A))" by (simp add: cSup_insert cong del: SUP_cong) lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ \(f ` A) \ \(g ` B)" using cInf_mono [of "g ` B" "f ` A"] by auto lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ \(f ` A) \ \(g ` B)" using cSup_mono [of "f ` A" "g ` B"] by auto lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ \(g ` B) \ \(f ` A)" by (rule cINF_mono) auto -lemma cSUP_subset_mono: "A \ {} \ bdd_above (g ` B) \ A \ B \ (\x. x \ B \ f x \ g x) \ \(f ` A) \ \(g ` B)" +lemma cSUP_subset_mono: + "\A \ {}; bdd_above (g ` B); A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (rule cSUP_mono) auto lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) lemma cINF_union: "A \ {} \ bdd_below (f ` A) \ B \ {} \ bdd_below (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un cong del: INF_cong) lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) lemma cSUP_union: "A \ {} \ bdd_above (f ` A) \ B \ {} \ bdd_above (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un cong del: SUP_cong) lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. inf (f a) (g a))" by (intro antisym le_infI cINF_greatest cINF_lower2) (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. sup (f a) (g a))" by (intro antisym le_supI cSUP_least cSUP_upper2) (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) lemma cInf_le_cSup: "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) end instance complete_lattice \ conditionally_complete_lattice by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) lemma cSup_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_bot}" assumes upper: "\x. x \ X \ x \ a" assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" proof cases assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cSup_eq_non_empty assms) lemma cInf_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" assumes upper: "\x. x \ X \ a \ x" assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" proof cases assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cInf_eq_non_empty assms) class conditionally_complete_linorder = conditionally_complete_lattice + linorder begin lemma less_cSup_iff: "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (\i\A. f i) < a \ (\x\A. f x < a)" using cInf_less_iff[of "f`A"] by auto lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (\i\A. f i) \ (\x\A. a < f x)" using less_cSup_iff[of "f`A"] by auto lemma less_cSupE: assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" by (metis cSup_least assms not_le that) lemma less_cSupD: "X \ {} \ z < Sup X \ \x\X. z < x" by (metis less_cSup_iff not_le_imp_less bdd_above_def) lemma cInf_lessD: "X \ {} \ Inf X < z \ \x\X. x < z" by (metis cInf_less_iff not_le_imp_less bdd_below_def) lemma complete_interval: assumes "a < b" and "P a" and "\ P b" shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" proof (rule exI [where x = "Sup {d. \x. a \ x \ x < d \ P x}"], auto) show "a \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule cSup_upper, auto simp: bdd_above_def) (metis \a < b\ \\ P b\ linear less_le) next show "Sup {d. \c. a \ c \ c < d \ P c} \ b" apply (rule cSup_least) apply auto apply (metis less_le_not_le) apply (metis \a \\ P b\ linear less_le) done next fix x assume x: "a \ x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" show "P x" apply (rule less_cSupE [OF lt], auto) apply (metis less_le_not_le) apply (metis x) done next fix d assume 0: "\x. a \ x \ x < d \ P x" thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule_tac cSup_upper, auto simp: bdd_above_def) (metis \a \\ P b\ linear less_le) qed end instance complete_linorder < conditionally_complete_linorder .. lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max) lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min) lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y" by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (simp add: cInf_eq_Min) lemma Sup_insert_finite: fixes S :: "'a::conditionally_complete_linorder set" shows "finite S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (simp add: cSup_insert sup_max) lemma finite_imp_less_Inf: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a < x\ \ a < Inf X" by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) lemma finite_less_Inf_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a < Inf X \ (\x \ X. a < x)" by (auto simp: cInf_eq_Min) lemma finite_imp_Sup_less: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a > x\ \ a > Sup X" by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) lemma finite_Sup_less_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a > Sup X \ (\x \ X. a > x)" by (auto simp: cSup_eq_Max) class linear_continuum = conditionally_complete_linorder + dense_linorder + assumes UNIV_not_singleton: "\a b::'a. a \ b" begin lemma ex_gt_or_lt: "\b. a < b \ b < a" by (metis UNIV_not_singleton neq_iff) end instantiation nat :: conditionally_complete_linorder begin definition "Sup (X::nat set) = Max X" definition "Inf (X::nat set) = (LEAST n. n \ X)" lemma bdd_above_nat: "bdd_above X \ finite (X::nat set)" proof assume "bdd_above X" then obtain z where "X \ {.. z}" by (auto simp: bdd_above_def) then show "finite X" by (rule finite_subset) simp qed simp instance proof fix x :: nat fix X :: "nat set" show "Inf X \ x" if "x \ X" "bdd_below X" using that by (simp add: Inf_nat_def Least_le) show "x \ Inf X" if "X \ {}" "\y. y \ X \ x \ y" using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) show "x \ Sup X" if "x \ X" "bdd_above X" using that by (simp add: Sup_nat_def bdd_above_nat) show "Sup X \ x" if "X \ {}" "\y. y \ X \ y \ x" proof - from that have "bdd_above X" by (auto simp: bdd_above_def) with that show ?thesis by (simp add: Sup_nat_def bdd_above_nat) qed qed end lemma Inf_nat_def1: fixes K::"nat set" assumes "K \ {}" shows "Inf K \ K" by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI) instantiation int :: conditionally_complete_linorder begin definition "Sup (X::int set) = (THE x. x \ X \ (\y\X. y \ x))" definition "Inf (X::int set) = - (Sup (uminus ` X))" instance proof { fix x :: int and X :: "int set" assume "X \ {}" "bdd_above X" then obtain x y where "X \ {..y}" "x \ X" by (auto simp: bdd_above_def) then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" by (auto simp: subset_eq) have "\!x\X. (\y\X. y \ x)" proof { fix z assume "z \ X" have "z \ Max (X \ {x..y})" proof cases assume "x \ z" with \z \ X\ \X \ {..y}\ *(1) show ?thesis by (auto intro!: Max_ge) next assume "\ x \ z" then have "z < x" by simp also have "x \ Max (X \ {x..y})" using \x \ X\ *(1) \x \ y\ by (intro Max_ge) auto finally show ?thesis by simp qed } note le = this with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto fix z assume *: "z \ X \ (\y\X. y \ z)" with le have "z \ Max (X \ {x..y})" by auto moreover have "Max (X \ {x..y}) \ z" using * ex by auto ultimately show "z = Max (X \ {x..y})" by auto qed then have "Sup X \ X \ (\y\X. y \ Sup X)" unfolding Sup_int_def by (rule theI') } note Sup_int = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_above X" then show "x \ Sup X" using Sup_int[of X] by auto } note le_Sup = this { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ y \ x" then show "Sup X \ x" using Sup_int[of X] by (auto simp: bdd_above_def) } note Sup_le = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_below X" then show "Inf X \ x" using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } qed end lemma interval_cases: fixes S :: "'a :: conditionally_complete_linorder set" assumes ivl: "\a b x. a \ S \ b \ S \ a \ x \ x \ b \ x \ S" shows "\a b. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" proof - define lower upper where "lower = {x. \s\S. s \ x}" and "upper = {x. \s\S. x \ s}" with ivl have "S = lower \ upper" by auto moreover have "\a. upper = UNIV \ upper = {} \ upper = {.. a} \ upper = {..< a}" proof cases assume *: "bdd_above S \ S \ {}" from * have "upper \ {.. Sup S}" by (auto simp: upper_def intro: cSup_upper2) moreover from * have "{..< Sup S} \ upper" by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) ultimately have "upper = {.. Sup S} \ upper = {..< Sup S}" unfolding ivl_disj_un(2)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_above S \ S \ {})" then have "upper = UNIV \ upper = {}" by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) then show ?thesis by auto qed moreover have "\b. lower = UNIV \ lower = {} \ lower = {b ..} \ lower = {b <..}" proof cases assume *: "bdd_below S \ S \ {}" from * have "lower \ {Inf S ..}" by (auto simp: lower_def intro: cInf_lower2) moreover from * have "{Inf S <..} \ lower" by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) ultimately have "lower = {Inf S ..} \ lower = {Inf S <..}" unfolding ivl_disj_un(1)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_below S \ S \ {})" then have "lower = UNIV \ lower = {}" by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) then show ?thesis by auto qed ultimately show ?thesis unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral) qed lemma cSUP_eq_cINF_D: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes eq: "(\x\A. f x) = (\x\A. f x)" and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" and a: "a \ A" shows "f a = (\x\A. f x)" apply (rule antisym) using a bdd apply (auto simp: cINF_lower) apply (metis eq cSUP_upper) done lemma cSUP_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_above (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_above (f ` B x)" using bdd_UN by (meson UN_upper bdd_above_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_above_def) then have bdd2: "bdd_above ((\x. \z\B x. f z) ` A)" unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\ z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) qed lemma cINF_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_below (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_below (f ` B x)" using bdd_UN by (meson UN_upper bdd_below_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_below_def) then have bdd2: "bdd_below ((\x. \z\B x. f z) ` A)" unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) ultimately show ?thesis by (rule order_antisym) qed lemma cSup_abs_le: fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set" shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Sup S\ \ a" apply (auto simp add: abs_le_iff intro: cSup_least) by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans) end diff --git a/src/HOL/Library/Equipollence.thy b/src/HOL/Library/Equipollence.thy --- a/src/HOL/Library/Equipollence.thy +++ b/src/HOL/Library/Equipollence.thy @@ -1,358 +1,696 @@ section \Equipollence and Other Relations Connected with Cardinality\ theory "Equipollence" - imports FuncSet + imports FuncSet begin subsection\Eqpoll\ definition eqpoll :: "'a set \ 'b set \ bool" (infixl "\" 50) where "eqpoll A B \ \f. bij_betw f A B" definition lepoll :: "'a set \ 'b set \ bool" (infixl "\" 50) where "lepoll A B \ \f. inj_on f A \ f ` A \ B" definition lesspoll :: "'a set \ 'b set \ bool" (infixl \\\ 50) where "A \ B == A \ B \ ~(A \ B)" lemma lepoll_empty_iff_empty [simp]: "A \ {} \ A = {}" by (auto simp: lepoll_def) lemma eqpoll_iff_card_of_ordIso: "A \ B \ ordIso2 (card_of A) (card_of B)" by (simp add: card_of_ordIso eqpoll_def) +lemma eqpoll_refl [iff]: "A \ A" + by (simp add: card_of_refl eqpoll_iff_card_of_ordIso) + lemma eqpoll_finite_iff: "A \ B \ finite A \ finite B" by (meson bij_betw_finite eqpoll_def) lemma eqpoll_iff_card: assumes "finite A" "finite B" shows "A \ B \ card A = card B" using assms by (auto simp: bij_betw_iff_card eqpoll_def) lemma lepoll_antisym: assumes "A \ B" "B \ A" shows "A \ B" using assms unfolding eqpoll_def lepoll_def by (metis Schroeder_Bernstein) lemma lepoll_trans [trans]: "\A \ B; B \ C\ \ A \ C" apply (clarsimp simp: lepoll_def) apply (rename_tac f g) apply (rule_tac x="g \ f" in exI) apply (auto simp: image_subset_iff inj_on_def) done lemma lepoll_trans1 [trans]: "\A \ B; B \ C\ \ A \ C" by (meson card_of_ordLeq eqpoll_iff_card_of_ordIso lepoll_def lepoll_trans ordIso_iff_ordLeq) lemma lepoll_trans2 [trans]: "\A \ B; B \ C\ \ A \ C" apply (clarsimp simp: eqpoll_def lepoll_def bij_betw_def) apply (rename_tac f g) apply (rule_tac x="g \ f" in exI) apply (auto simp: image_subset_iff inj_on_def) done lemma eqpoll_sym: "A \ B \ B \ A" unfolding eqpoll_def using bij_betw_the_inv_into by auto lemma eqpoll_trans [trans]: "\A \ B; B \ C\ \ A \ C" unfolding eqpoll_def using bij_betw_trans by blast lemma eqpoll_imp_lepoll: "A \ B \ A \ B" unfolding eqpoll_def lepoll_def by (metis bij_betw_def order_refl) lemma subset_imp_lepoll: "A \ B \ A \ B" by (force simp: lepoll_def) +lemma lepoll_refl [iff]: "A \ A" + by (simp add: subset_imp_lepoll) + lemma lepoll_iff: "A \ B \ (\g. A \ g ` B)" unfolding lepoll_def proof safe fix g assume "A \ g ` B" then show "\f. inj_on f A \ f ` A \ B" by (rule_tac x="inv_into B g" in exI) (auto simp: inv_into_into inj_on_inv_into) qed (metis image_mono the_inv_into_onto) +lemma empty_lepoll [iff]: "{} \ A" + by (simp add: lepoll_iff) + lemma subset_image_lepoll: "B \ f ` A \ B \ A" by (auto simp: lepoll_iff) lemma image_lepoll: "f ` A \ A" by (auto simp: lepoll_iff) lemma infinite_le_lepoll: "infinite A \ (UNIV::nat set) \ A" apply (auto simp: lepoll_def) apply (simp add: infinite_countable_subset) using infinite_iff_countable_subset by blast +lemma lepoll_Pow_self: "A \ Pow A" + unfolding lepoll_def inj_def + proof (intro exI conjI) + show "inj_on (\x. {x}) A" + by (auto simp: inj_on_def) +qed auto + lemma bij_betw_iff_bijections: "bij_betw f A B \ (\g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs apply (rule_tac x="the_inv_into A f" in exI) apply (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into) done next assume ?rhs then show ?lhs by (auto simp: bij_betw_def inj_on_def image_def; metis) qed lemma eqpoll_iff_bijections: "A \ B \ (\f g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" by (auto simp: eqpoll_def bij_betw_iff_bijections) lemma lepoll_restricted_funspace: "{f. f ` A \ B \ {x. f x \ k x} \ A \ finite {x. f x \ k x}} \ Fpow (A \ B)" proof - have *: "\U \ Fpow (A \ B). f = (\x. if \y. (x, y) \ U then SOME y. (x,y) \ U else k x)" if "f ` A \ B" "{x. f x \ k x} \ A" "finite {x. f x \ k x}" for f apply (rule_tac x="(\x. (x, f x)) ` {x. f x \ k x}" in bexI) using that by (auto simp: image_def Fpow_def) show ?thesis apply (rule subset_image_lepoll [where f = "\U x. if \y. (x,y) \ U then @y. (x,y) \ U else k x"]) using * by (auto simp: image_def) qed lemma singleton_lepoll: "{x} \ insert y A" by (force simp: lepoll_def) lemma singleton_eqpoll: "{x} \ {y}" by (blast intro: lepoll_antisym singleton_lepoll) lemma subset_singleton_iff_lepoll: "(\x. S \ {x}) \ S \ {()}" proof safe show "S \ {()}" if "S \ {x}" for x using subset_imp_lepoll [OF that] by (simp add: singleton_eqpoll lepoll_trans2) show "\x. S \ {x}" if "S \ {()}" by (metis (no_types, hide_lams) image_empty image_insert lepoll_iff that) qed +lemma infinite_insert_lepoll: + assumes "infinite A" shows "insert a A \ A" +proof - + obtain f :: "nat \ 'a" where "inj f" and f: "range f \ A" + using assms infinite_countable_subset by blast + let ?g = "(\z. if z=a then f 0 else if z \ range f then f (Suc (inv f z)) else z)" + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on ?g (insert a A)" + using inj_on_eq_iff [OF \inj f\] + by (auto simp: inj_on_def) + show "?g ` insert a A \ A" + using f by auto + qed +qed + +lemma infinite_insert_eqpoll: "infinite A \ insert a A \ A" + by (simp add: lepoll_antisym infinite_insert_lepoll subset_imp_lepoll subset_insertI) + +lemma finite_lepoll_infinite: + assumes "infinite A" "finite B" shows "B \ A" +proof - + have "B \ (UNIV::nat set)" + unfolding lepoll_def + using finite_imp_inj_to_nat_seg [OF \finite B\] by blast + then show ?thesis + using \infinite A\ infinite_le_lepoll lepoll_trans by auto +qed subsection\The strict relation\ lemma lesspoll_not_refl [iff]: "~ (i \ i)" by (simp add: lepoll_antisym lesspoll_def) lemma lesspoll_imp_lepoll: "A \ B ==> A \ B" by (unfold lesspoll_def, blast) lemma lepoll_iff_leqpoll: "A \ B \ A \ B | A \ B" using eqpoll_imp_lepoll lesspoll_def by blast lemma lesspoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (meson eqpoll_sym lepoll_antisym lepoll_trans lepoll_trans1 lesspoll_def) lemma lesspoll_trans1 [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (meson eqpoll_sym lepoll_antisym lepoll_trans lepoll_trans1 lesspoll_def) lemma lesspoll_trans2 [trans]: "\X \ Y; Y \ Z\ \ X \ Z" by (meson eqpoll_imp_lepoll eqpoll_sym lepoll_antisym lepoll_trans lesspoll_def) lemma eq_lesspoll_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" using eqpoll_imp_lepoll lesspoll_trans1 by blast lemma lesspoll_eq_trans [trans]: "\X \ Y; Y \ Z\ \ X \ Z" using eqpoll_imp_lepoll lesspoll_trans2 by blast -subsection\Cartesian products\ +lemma lesspoll_Pow_self: "A \ Pow A" + unfolding lesspoll_def bij_betw_def eqpoll_def + by (meson lepoll_Pow_self Cantors_paradox) + +lemma finite_lesspoll_infinite: + assumes "infinite A" "finite B" shows "B \ A" + by (meson assms eqpoll_finite_iff finite_lepoll_infinite lesspoll_def) + +subsection\Mapping by an injection\ + +lemma inj_on_image_eqpoll_self: "inj_on f A \ f ` A \ A" + by (meson bij_betw_def eqpoll_def eqpoll_sym) + +lemma inj_on_image_lepoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (meson assms image_lepoll lepoll_def lepoll_trans order_refl) + +lemma inj_on_image_lepoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (meson assms eq_iff image_lepoll lepoll_def lepoll_trans) + +lemma inj_on_image_lesspoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (meson assms image_lepoll le_less lepoll_def lesspoll_trans1) + +lemma inj_on_image_lesspoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (meson assms eqpoll_sym inj_on_image_eqpoll_self lesspoll_eq_trans) + +lemma inj_on_image_eqpoll_1 [simp]: + assumes "inj_on f A" shows "f ` A \ B \ A \ B" + by (metis assms eqpoll_trans inj_on_image_eqpoll_self eqpoll_sym) + +lemma inj_on_image_eqpoll_2 [simp]: + assumes "inj_on f B" shows "A \ f ` B \ A \ B" + by (metis assms inj_on_image_eqpoll_1 eqpoll_sym) + +subsection \Inserting elements into sets\ + +lemma insert_lepoll_insertD: + assumes "insert u A \ insert v B" "u \ A" "v \ B" shows "A \ B" +proof - + obtain f where inj: "inj_on f (insert u A)" and fim: "f ` (insert u A) \ insert v B" + by (meson assms lepoll_def) + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + let ?g = "\x\A. if f x = v then f u else f x" + show "inj_on ?g A" + using inj \u \ A\ by (auto simp: inj_on_def) + show "?g ` A \ B" + using fim \u \ A\ image_subset_iff inj inj_on_image_mem_iff by fastforce + qed +qed + +lemma insert_eqpoll_insertD: "\insert u A \ insert v B; u \ A; v \ B\ \ A \ B" + by (meson insert_lepoll_insertD eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) + +lemma insert_lepoll_cong: + assumes "A \ B" "b \ B" shows "insert a A \ insert b B" +proof - + obtain f where f: "inj_on f A" "f ` A \ B" + by (meson assms lepoll_def) + let ?f = "\u \ insert a A. if u=a then b else f u" + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on ?f (insert a A)" + using f \b \ B\ by (auto simp: inj_on_def) + show "?f ` insert a A \ insert b B" + using f \b \ B\ by auto + qed +qed + +lemma insert_eqpoll_cong: + "\A \ B; a \ A; b \ B\ \ insert a A \ insert b B" + apply (rule lepoll_antisym) + apply (simp add: eqpoll_imp_lepoll insert_lepoll_cong)+ + by (meson eqpoll_imp_lepoll eqpoll_sym insert_lepoll_cong) + +lemma insert_eqpoll_insert_iff: + "\a \ A; b \ B\ \ insert a A \ insert b B \ A \ B" + by (meson insert_eqpoll_insertD insert_eqpoll_cong) + +lemma insert_lepoll_insert_iff: + " \a \ A; b \ B\ \ (insert a A \ insert b B) \ (A \ B)" + by (meson insert_lepoll_insertD insert_lepoll_cong) + +lemma less_imp_insert_lepoll: + assumes "A \ B" shows "insert a A \ B" +proof - + obtain f where "inj_on f A" "f ` A \ B" + using assms by (metis bij_betw_def eqpoll_def lepoll_def lesspoll_def psubset_eq) + then obtain b where b: "b \ B" "b \ f ` A" + by auto + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + show "inj_on (f(a:=b)) (insert a A)" + using b \inj_on f A\ by (auto simp: inj_on_def) + show "(f(a:=b)) ` insert a A \ B" + using \f ` A \ B\ by (auto simp: b) + qed +qed + +lemma finite_insert_lepoll: "finite A \ (insert a A \ A) \ (a \ A)" +proof (induction A rule: finite_induct) + case (insert x A) + then show ?case + apply (auto simp: insert_absorb) + by (metis insert_commute insert_iff insert_lepoll_insertD) +qed auto + + +subsection\Binary sums and unions\ + +lemma Un_lepoll_mono: + assumes "A \ C" "B \ D" "disjnt C D" shows "A \ B \ C \ D" +proof - + obtain f g where inj: "inj_on f A" "inj_on g B" and fg: "f ` A \ C" "g ` B \ D" + by (meson assms lepoll_def) + have "inj_on (\x. if x \ A then f x else g x) (A \ B)" + using inj \disjnt C D\ fg unfolding disjnt_iff + by (fastforce intro: inj_onI dest: inj_on_contraD split: if_split_asm) + with fg show ?thesis + unfolding lepoll_def + by (rule_tac x="\x. if x \ A then f x else g x" in exI) auto +qed + +lemma Un_eqpoll_cong: "\A \ C; B \ D; disjnt A B; disjnt C D\ \ A \ B \ C \ D" + by (meson Un_lepoll_mono eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) + +lemma sum_lepoll_mono: + assumes "A \ C" "B \ D" shows "A <+> B \ C <+> D" +proof - + obtain f g where "inj_on f A" "f ` A \ C" "inj_on g B" "g ` B \ D" + by (meson assms lepoll_def) + then show ?thesis + unfolding lepoll_def + by (rule_tac x="case_sum (Inl \ f) (Inr \ g)" in exI) (force simp: inj_on_def) +qed + +lemma sum_eqpoll_cong: "\A \ C; B \ D\ \ A <+> B \ C <+> D" + by (meson eqpoll_imp_lepoll eqpoll_sym lepoll_antisym sum_lepoll_mono) + +subsection\Binary Cartesian products\ + +lemma times_square_lepoll: "A \ A \ A" + unfolding lepoll_def inj_def +proof (intro exI conjI) + show "inj_on (\x. (x,x)) A" + by (auto simp: inj_on_def) +qed auto + +lemma times_commute_eqpoll: "A \ B \ B \ A" + unfolding eqpoll_def + by (force intro: bij_betw_byWitness [where f = "\(x,y). (y,x)" and f' = "\(x,y). (y,x)"]) + +lemma times_assoc_eqpoll: "(A \ B) \ C \ A \ (B \ C)" + unfolding eqpoll_def + by (force intro: bij_betw_byWitness [where f = "\((x,y),z). (x,(y,z))" and f' = "\(x,(y,z)). ((x,y),z)"]) + +lemma times_singleton_eqpoll: "{a} \ A \ A" +proof - + have "{a} \ A = (\x. (a,x)) ` A" + by auto + also have "\ \ A" + proof (rule inj_on_image_eqpoll_self) + show "inj_on (Pair a) A" + by (auto simp: inj_on_def) + qed + finally show ?thesis . +qed + +lemma times_lepoll_mono: + assumes "A \ C" "B \ D" shows "A \ B \ C \ D" +proof - + obtain f g where "inj_on f A" "f ` A \ C" "inj_on g B" "g ` B \ D" + by (meson assms lepoll_def) + then show ?thesis + unfolding lepoll_def + by (rule_tac x="\(x,y). (f x, g y)" in exI) (auto simp: inj_on_def) +qed + +lemma times_eqpoll_cong: "\A \ C; B \ D\ \ A \ B \ C \ D" + by (metis eqpoll_imp_lepoll eqpoll_sym lepoll_antisym times_lepoll_mono) + +lemma + assumes "B \ {}" shows lepoll_times1: "A \ A \ B" and lepoll_times2: "A \ B \ A" + using assms lepoll_iff by fastforce+ + +lemma times_0_eqpoll: "{} \ A \ {}" + by (simp add: eqpoll_iff_bijections) + +lemma Sigma_lepoll_mono: + assumes "A \ C" "\x. x \ A \ B x \ D x" shows "Sigma A B \ Sigma C D" +proof - + have "\x. x \ A \ \f. inj_on f (B x) \ f ` (B x) \ D x" + by (meson assms lepoll_def) + then obtain f where "\x. x \ A \ inj_on (f x) (B x) \ f x ` B x \ D x" + by metis + with \A \ C\ show ?thesis + unfolding lepoll_def inj_on_def + by (rule_tac x="\(x,y). (x, f x y)" in exI) force +qed + +lemma sum_times_distrib_eqpoll: "(A <+> B) \ C \ (A \ C) <+> (B \ C)" + unfolding eqpoll_def +proof + show "bij_betw (\(x,z). case_sum(\y. Inl(y,z)) (\y. Inr(y,z)) x) ((A <+> B) \ C) (A \ C <+> B \ C)" + by (rule bij_betw_byWitness [where f' = "case_sum (\(x,z). (Inl x, z)) (\(y,z). (Inr y, z))"]) auto +qed + +lemma prod_insert_eqpoll: + assumes "a \ A" shows "insert a A \ B \ B <+> A \ B" + unfolding eqpoll_def + proof + show "bij_betw (\(x,y). if x=a then Inl y else Inr (x,y)) (insert a A \ B) (B <+> A \ B)" + by (rule bij_betw_byWitness [where f' = "case_sum (\y. (a,y)) id"]) (auto simp: assms) +qed + +subsection\General Unions\ + +lemma Union_eqpoll_Times: + assumes B: "\x. x \ A \ F x \ B" and disj: "pairwise (\x y. disjnt (F x) (F y)) A" + shows "(\x\A. F x) \ A \ B" +proof (rule lepoll_antisym) + obtain b where b: "\x. x \ A \ bij_betw (b x) (F x) B" + using B unfolding eqpoll_def by metis + show "\(F ` A) \ A \ B" + unfolding lepoll_def + proof (intro exI conjI) + define \ where "\ \ \z. THE x. x \ A \ z \ F x" + have \: "\ z = x" if "x \ A" "z \ F x" for x z + unfolding \_def + apply (rule the_equality) + apply (simp add: that) + by (metis disj disjnt_iff pairwiseD that) + let ?f = "\z. (\ z, b (\ z) z)" + show "inj_on ?f (\(F ` A))" + unfolding inj_on_def + by clarify (metis \ b bij_betw_inv_into_left) + show "?f ` \(F ` A) \ A \ B" + using \ b bij_betwE by blast + qed + show "A \ B \ \(F ` A)" + unfolding lepoll_def + proof (intro exI conjI) + let ?f = "\(x,y). inv_into (F x) (b x) y" + have *: "inv_into (F x) (b x) y \ F x" if "x \ A" "y \ B" for x y + by (metis b bij_betw_imp_surj_on inv_into_into that) + then show "inj_on ?f (A \ B)" + unfolding inj_on_def + by clarsimp (metis (mono_tags, lifting) b bij_betw_inv_into_right disj disjnt_iff pairwiseD) + show "?f ` (A \ B) \ \ (F ` A)" + by clarsimp (metis b bij_betw_imp_surj_on inv_into_into) + qed +qed + +lemma UN_lepoll_UN: + assumes A: "\x. x \ A \ B x \ C x" + and disj: "pairwise (\x y. disjnt (C x) (C y)) A" + shows "\ (B`A) \ \ (C`A)" +proof - + obtain f where f: "\x. x \ A \ inj_on (f x) (B x) \ f x ` (B x) \ (C x)" + using A unfolding lepoll_def by metis + show ?thesis + unfolding lepoll_def + proof (intro exI conjI) + define \ where "\ \ \z. @x. x \ A \ z \ B x" + have \: "\ z \ A \ z \ B (\ z)" if "x \ A" "z \ B x" for x z + unfolding \_def by (metis (mono_tags, lifting) someI_ex that) + let ?f = "\z. (f (\ z) z)" + show "inj_on ?f (\(B ` A))" + using disj f unfolding inj_on_def disjnt_iff pairwise_def image_subset_iff + by (metis UN_iff \) + show "?f ` \ (B ` A) \ \ (C ` A)" + using \ f unfolding image_subset_iff by blast + qed +qed + +lemma UN_eqpoll_UN: + assumes A: "\x. x \ A \ B x \ C x" + and B: "pairwise (\x y. disjnt (B x) (B y)) A" + and C: "pairwise (\x y. disjnt (C x) (C y)) A" + shows "(\x\A. B x) \ (\x\A. C x)" +proof (rule lepoll_antisym) + show "\ (B ` A) \ \ (C ` A)" + by (meson A C UN_lepoll_UN eqpoll_imp_lepoll) + show "\ (C ` A) \ \ (B ` A)" + by (simp add: A B UN_lepoll_UN eqpoll_imp_lepoll eqpoll_sym) +qed + +subsection\General Cartesian products (Pi)\ lemma PiE_sing_eqpoll_self: "({a} \\<^sub>E B) \ B" proof - have 1: "x = y" if "x \ {a} \\<^sub>E B" "y \ {a} \\<^sub>E B" "x a = y a" for x y by (metis IntD2 PiE_def extensionalityI singletonD that) have 2: "x \ (\h. h a) ` ({a} \\<^sub>E B)" if "x \ B" for x using that by (rule_tac x="\z\{a}. x" in image_eqI) auto show ?thesis unfolding eqpoll_def bij_betw_def inj_on_def by (force intro: 1 2) qed lemma lepoll_funcset_right: "B \ B' \ A \\<^sub>E B \ A \\<^sub>E B'" apply (auto simp: lepoll_def inj_on_def) apply (rule_tac x = "\g. \z \ A. f(g z)" in exI) apply (auto simp: fun_eq_iff) apply (metis PiE_E) by blast lemma lepoll_funcset_left: assumes "B \ {}" "A \ A'" shows "A \\<^sub>E B \ A' \\<^sub>E B" proof - obtain b where "b \ B" using assms by blast obtain f where "inj_on f A" and fim: "f ` A \ A'" using assms by (auto simp: lepoll_def) then obtain h where h: "\x. x \ A \ h (f x) = x" using the_inv_into_f_f by fastforce let ?F = "\g. \u \ A'. if h u \ A then g(h u) else b" show ?thesis unfolding lepoll_def inj_on_def proof (intro exI conjI ballI impI ext) fix k l x assume k: "k \ A \\<^sub>E B" and l: "l \ A \\<^sub>E B" and "?F k = ?F l" then have "?F k (f x) = ?F l (f x)" by simp then show "k x = l x" apply (auto simp: h split: if_split_asm) apply (metis PiE_arb h k l) apply (metis (full_types) PiE_E h k l) using fim k l by fastforce next show "?F ` (A \\<^sub>E B) \ A' \\<^sub>E B" using \b \ B\ by force qed qed lemma lepoll_funcset: "\B \ {}; A \ A'; B \ B'\ \ A \\<^sub>E B \ A' \\<^sub>E B'" by (rule lepoll_trans [OF lepoll_funcset_right lepoll_funcset_left]) auto lemma lepoll_PiE: assumes "\i. i \ A \ B i \ C i" shows "PiE A B \ PiE A C" proof - obtain f where f: "\i. i \ A \ inj_on (f i) (B i) \ (f i) ` B i \ C i" using assms unfolding lepoll_def by metis then show ?thesis unfolding lepoll_def apply (rule_tac x = "\g. \i \ A. f i (g i)" in exI) apply (auto simp: inj_on_def) apply (rule PiE_ext, auto) apply (metis (full_types) PiE_mem restrict_apply') by blast qed lemma card_le_PiE_subindex: assumes "A \ A'" "Pi\<^sub>E A' B \ {}" shows "PiE A B \ PiE A' B" proof - have "\x. x \ A' \ \y. y \ B x" using assms by blast then obtain g where g: "\x. x \ A' \ g x \ B x" by metis let ?F = "\f x. if x \ A then f x else if x \ A' then g x else undefined" have "Pi\<^sub>E A B \ (\f. restrict f A) ` Pi\<^sub>E A' B" proof show "f \ Pi\<^sub>E A B \ f \ (\f. restrict f A) ` Pi\<^sub>E A' B" for f using \A \ A'\ by (rule_tac x="?F f" in image_eqI) (auto simp: g fun_eq_iff) qed then have "Pi\<^sub>E A B \ (\f. \i \ A. f i) ` Pi\<^sub>E A' B" by (simp add: subset_imp_lepoll) also have "\ \ PiE A' B" by (rule image_lepoll) finally show ?thesis . qed lemma finite_restricted_funspace: assumes "finite A" "finite B" shows "finite {f. f ` A \ B \ {x. f x \ k x} \ A}" (is "finite ?F") proof (rule finite_subset) show "finite ((\U x. if \y. (x,y) \ U then @y. (x,y) \ U else k x) ` Pow(A \ B))" (is "finite ?G") using assms by auto show "?F \ ?G" proof fix f assume "f \ ?F" then show "f \ ?G" by (rule_tac x="(\x. (x,f x)) ` {x. f x \ k x}" in image_eqI) (auto simp: fun_eq_iff image_def) qed qed proposition finite_PiE_iff: "finite(PiE I S) \ PiE I S = {} \ finite {i \ I. ~(\a. S i \ {a})} \ (\i \ I. finite(S i))" (is "?lhs = ?rhs") proof (cases "PiE I S = {}") case False define J where "J \ {i \ I. \a. S i \ {a}}" show ?thesis proof assume L: ?lhs have "infinite (Pi\<^sub>E I S)" if "infinite J" proof - have "(UNIV::nat set) \ (UNIV::(nat\bool) set)" proof - have "\N::nat set. inj_on (=) N" by (simp add: inj_on_def) then show ?thesis by (meson infinite_iff_countable_subset infinite_le_lepoll top.extremum) qed also have "\ = (UNIV::nat set) \\<^sub>E (UNIV::bool set)" by auto also have "\ \ J \\<^sub>E (UNIV::bool set)" apply (rule lepoll_funcset_left) using infinite_le_lepoll that by auto also have "\ \ Pi\<^sub>E J S" proof - have *: "(UNIV::bool set) \ S i" if "i \ I" and "\a. \ S i \ {a}" for i proof - obtain a b where "{a,b} \ S i" "a \ b" by (metis \\a. \ S i \ {a}\ all_not_in_conv empty_subsetI insertCI insert_subset set_eq_subset subsetI) then show ?thesis apply (clarsimp simp: lepoll_def inj_on_def) apply (rule_tac x="\x. if x then a else b" in exI, auto) done qed show ?thesis by (auto simp: * J_def intro: lepoll_PiE) qed also have "\ \ Pi\<^sub>E I S" using False by (auto simp: J_def intro: card_le_PiE_subindex) finally have "(UNIV::nat set) \ Pi\<^sub>E I S" . then show ?thesis by (simp add: infinite_le_lepoll) qed moreover have "finite (S i)" if "i \ I" for i proof (rule finite_subset) obtain f where f: "f \ PiE I S" using False by blast show "S i \ (\f. f i) ` Pi\<^sub>E I S" proof show "s \ (\f. f i) ` Pi\<^sub>E I S" if "s \ S i" for s using that f \i \ I\ by (rule_tac x="\j. if j = i then s else f j" in image_eqI) auto qed next show "finite ((\x. x i) ` Pi\<^sub>E I S)" using L by blast qed ultimately show ?rhs using L by (auto simp: J_def False) next assume R: ?rhs have "\i \ I - J. \a. S i = {a}" using False J_def by blast then obtain a where a: "\i \ I - J. S i = {a i}" by metis let ?F = "{f. f ` J \ (\i \ J. S i) \ {i. f i \ (if i \ I then a i else undefined)} \ J}" have *: "finite (Pi\<^sub>E I S)" if "finite J" and "\i\I. finite (S i)" proof (rule finite_subset) show "Pi\<^sub>E I S \ ?F" apply safe using J_def apply blast by (metis DiffI PiE_E a singletonD) show "finite ?F" proof (rule finite_restricted_funspace [OF \finite J\]) show "finite (\ (S ` J))" using that J_def by blast qed qed show ?lhs using R by (auto simp: * J_def) qed qed auto corollary finite_funcset_iff: "finite(I \\<^sub>E S) \ (\a. S \ {a}) \ I = {} \ finite I \ finite S" apply (auto simp: finite_PiE_iff PiE_eq_empty_iff dest: not_finite_existsD) using finite.simps by auto end