# HG changeset patch # User wenzelm # Date 1545221222 -3600 # Node ID 3b89c6b723a29ca99ba8313ad0a40feb176c0352 # Parent 4880575ec8a17a950b17e85b75a3cbe9021e0968# Parent 2cc9212342a728cc77ac6a8e67b5e5b8de9f23a9 merged diff -r 2cc9212342a7 -r 3b89c6b723a2 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Wed Dec 19 12:26:38 2018 +0100 +++ b/src/HOL/Hilbert_Choice.thy Wed Dec 19 13:07:02 2018 +0100 @@ -909,13 +909,14 @@ context complete_distrib_lattice begin -lemma Sup_Inf: "Sup (Inf ` A) = Inf (Sup ` {f ` A | f . (\ Y \ A . f Y \ Y)})" + +lemma Sup_Inf: "\ (Inf ` A) = \ (Sup ` {f ` A |f. \B\A. f B \ B})" proof (rule antisym) - show "\(Inf ` A) \ \(Sup ` {f ` A |f. \Y\A. f Y \ Y})" + show "\ (Inf ` A) \ \ (Sup ` {f ` A |f. \B\A. f B \ B})" apply (rule Sup_least, rule INF_greatest) using Inf_lower2 Sup_upper by auto next - show "\(Sup ` {f ` A |f. \Y\A. f Y \ Y}) \ \(Inf ` A)" + show "\ (Sup ` {f ` A |f. \B\A. f B \ B}) \ \ (Inf ` A)" proof (simp add: Inf_Sup, rule SUP_least, simp, safe) fix f assume "\Y. (\f. Y = f ` A \ (\Y\A. f Y \ Y)) \ f Y \ Y" @@ -976,7 +977,7 @@ using dual_complete_distrib_lattice by (rule complete_distrib_lattice.sup_Inf) -lemma INF_SUP: "(INF y. SUP x. ((P x y)::'a)) = (SUP x. INF y. P (x y) y)" +lemma INF_SUP: "(\y. \x. P x y) = (\f. \x. P (f x) x)" proof (rule antisym) show "(SUP x. INF y. P (x y) y) \ (INF y. SUP x. P x y)" by (rule SUP_least, rule INF_greatest, rule SUP_upper2, simp_all, rule INF_lower2, simp, blast) @@ -1018,30 +1019,33 @@ by simp qed -lemma INF_SUP_set: "(\x\A. \(g ` x)) = (\x\{f ` A |f. \Y\A. f Y \ Y}. \(g ` x))" +lemma INF_SUP_set: "(\B\A. \(g ` B)) = (\B\{f ` A |f. \C\A. f C \ C}. \(g ` B))" proof (rule antisym) - have [simp]: "\f xa. \Y\A. f Y \ Y \ xa \ A \ (\x\A. g (f x)) \ g (f xa)" - by (rule INF_lower2, blast+) - have B: "\f xa. \Y\A. f Y \ Y \ xa \ A \ f xa \ xa" - by blast - have A: "\f xa. \Y\A. f Y \ Y \ xa \ A \ (\x\A. g (f x)) \ \(g ` xa)" - by (rule SUP_upper2, rule B, simp_all, simp) - show "(\x\{f ` A |f. \Y\A. f Y \ Y}. \a\x. g a) \ (\x\A. \a\x. g a)" - apply (rule SUP_least, simp, safe, rule INF_greatest, simp) - by (rule A) + have "\ ((g \ f) ` A) \ \ (g ` B)" if "\B. B \ A \ f B \ B" and "B \ A" + for f and B + using that by (auto intro: SUP_upper2 INF_lower2) + then show "(\x\{f ` A |f. \Y\A. f Y \ Y}. \a\x. g a) \ (\x\A. \a\x. g a)" + by (auto intro!: SUP_least INF_greatest) next show "(\x\A. \a\x. g a) \ (\x\{f ` A |f. \Y\A. f Y \ Y}. \a\x. g a)" proof (cases "{} \ A") case True then show ?thesis - by (rule INF_lower2, simp_all) + by (rule INF_lower2) simp_all next case False - have [simp]: "\x xa xb. xb \ A \ x xb \ xb \ (\xa. if xa \ A then if x xa \ xa then g (x xa) else \ else \) \ g (x xb)" + have *: "\f B. B \ A \ f B \ B \ + (\B. if B \ A then if f B \ B then g (f B) else \ else \) \ g (f B)" + by (rule INF_lower2, auto) + have **: "\f B. B \ A \ f B \ B \ + (\B. if B \ A then if f B \ B then g (f B) else \ else \) \ g (SOME x. x \ B)" by (rule INF_lower2, auto) - have [simp]: " \x xa y. y \ A \ x y \ y \ (\xa. if xa \ A then if x xa \ xa then g (x xa) else \ else \) \ g (SOME x. x \ y)" - by (rule INF_lower2, auto) - have [simp]: "\x. (\xa. if xa \ A then if x xa \ xa then g (x xa) else \ else \) \ (\x\{f ` A |f. \Y\A. f Y \ Y}. \x\x. g x)" + have ****: "\f B. B \ A \ + (\B. if B \ A then if f B \ B then g (f B) else \ else \) + \ (if f B \ B then g (f B) else g (SOME x. x \ B))" + by (rule INF_lower2) auto + have ***: "\x. (\B. if B \ A then if x B \ B then g (x B) else \ else \) + \ (\x\{f ` A |f. \Y\A. f Y \ Y}. \x\x. g x)" proof - fix x define F where "F = (\ (y::'b set) . if x y \ y then x y else (SOME x . x \y))" @@ -1051,7 +1055,10 @@ using B by blast show "(\xa. if xa \ A then if x xa \ xa then g (x xa) else \ else \) \ (\x\{f ` A |f. \Y\A. f Y \ Y}. \x\x. g x)" using A apply (rule SUP_upper2) - by (simp add: F_def, rule INF_greatest, auto) + apply (simp add: F_def) + apply (rule INF_greatest) + apply (auto simp add: * **) + done qed {fix x @@ -1071,16 +1078,16 @@ also have "... = (\x. \xa. if xa \ A then if x xa \ xa then g (x xa) else \ else \)" by (simp add: INF_SUP) also have "... \ (\x\{f ` A |f. \Y\A. f Y \ Y}. \a\x. g a)" - by (rule SUP_least, simp) + by (rule SUP_least, simp add: ***) finally show ?thesis by simp qed qed -lemma SUP_INF: "(SUP y. INF x. ((P x y)::'a)) = (INF x. SUP y. P (x y) y)" +lemma SUP_INF: "(\y. \x. P x y) = (\x. \y. P (x y) y)" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP) -lemma SUP_INF_set: "(\x\A. \(g ` x)) = (\x\{f ` A |f. \Y\A. f Y \ Y}. \(g ` x))" +lemma SUP_INF_set: "(\x\A. \ (g ` x)) = (\x\{f ` A |f. \Y\A. f Y \ Y}. \ (g ` x))" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP_set) @@ -1096,7 +1103,6 @@ lemma inf_SUP: "a \ (\b\B. f b) = (\b\B. a \ f b)" by (simp add: inf_Sup) - lemma Inf_sup: "\B \ a = (\b\B. b \ a)" by (simp add: sup_Inf sup_commute)