# HG changeset patch # User haftmann # Date 1312478992 -7200 # Node ID cb768f4ceaf9074c2b5bc26fe6d95fbb0e4356cd # Parent ce4e3090f01a9742010182508d3e6963be5bc6f1 solving duality problem for complete_distrib_lattice; tuned diff -r ce4e3090f01a -r cb768f4ceaf9 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Thu Aug 04 07:33:08 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Thu Aug 04 19:29:52 2011 +0200 @@ -184,10 +184,10 @@ qed definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - INF_def: "INFI A f = \ (f ` A)" + INF_def: "INFI A f = \(f ` A)" definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - SUP_def: "SUPR A f = \ (f ` A)" + SUP_def: "SUPR A f = \(f ` A)" text {* Note: must use names @{const INFI} and @{const SUPR} here instead of @@ -390,6 +390,26 @@ "(\b. A b) = A True \ A False" by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute) +lemma INF_foundation_dual [no_atp]: + "complete_lattice.SUPR Inf = INFI" +proof (rule ext)+ + interpret dual: complete_lattice Sup Inf "op \" "op >" sup inf \ \ + by (fact dual_complete_lattice) + fix f :: "'b \ 'a" and A + show "complete_lattice.SUPR Inf A f = (\a\A. f a)" + by (simp only: dual.SUP_def INF_def) +qed + +lemma SUP_foundation_dual [no_atp]: + "complete_lattice.INFI Sup = SUPR" +proof (rule ext)+ + interpret dual: complete_lattice Sup Inf "op \" "op >" sup inf \ \ + by (fact dual_complete_lattice) + fix f :: "'b \ 'a" and A + show "complete_lattice.INFI Sup A f = (\a\A. f a)" + by (simp only: dual.INF_def SUP_def) +qed + end class complete_distrib_lattice = complete_lattice + @@ -397,12 +417,13 @@ assumes sup_Inf: "a \ \B = (\b\B. a \ b)" begin -(*lemma dual_complete_distrib_lattice: +lemma dual_complete_distrib_lattice: "class.complete_distrib_lattice Sup Inf (op \) (op >) sup inf \ \" apply (rule class.complete_distrib_lattice.intro) apply (fact dual_complete_lattice) apply (rule class.complete_distrib_lattice_axioms.intro) - apply (simp_all add: inf_Sup sup_Inf)*) + apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf) + done subclass distrib_lattice proof -- {* Question: is it sufficient to include @{class distrib_lattice} and prove @{text inf_Sup} and @{text sup_Inf} from that? *} @@ -413,13 +434,12 @@ end -class complete_boolean_algebra = boolean_algebra + complete_lattice -- {* Question: is this - also a @{class complete_distrib_lattice}? *} +class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice begin lemma dual_complete_boolean_algebra: "class.complete_boolean_algebra Sup Inf (op \) (op >) sup inf \ \ (\x y. x \ - y) uminus" - by (rule class.complete_boolean_algebra.intro, rule dual_complete_lattice, rule dual_boolean_algebra) + by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) lemma uminus_Inf: "- (\A) = \(uminus ` A)" @@ -543,7 +563,7 @@ by (auto simp add: Bex_def SUP_def Sup_bool_def) qed -instance bool :: "{complete_distrib_lattice, complete_boolean_algebra}" proof +instance bool :: complete_boolean_algebra proof qed (auto simp add: Inf_bool_def Sup_bool_def) instantiation "fun" :: (type, complete_lattice) complete_lattice @@ -1024,7 +1044,7 @@ subsection {* Distributive laws *} lemma Int_Union: "A \ \B = (\C\B. A \ C)" - by blast + by (fact inf_Sup) lemma Int_Union2: "\B \ A = (\C\B. C \ A)" by blast @@ -1039,7 +1059,7 @@ by blast lemma Un_Inter: "A \ \B = (\C\B. A \ C)" - by blast + by (fact sup_Inf) lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" by blast @@ -1078,9 +1098,9 @@ lemma UN_simps [simp]: "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" - "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" + "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" "\A B C. (\x\C. A \ B x) = ((if C={} then {} else A \ (\x\C. B x)))" - "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \B)" + "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" "\A B C. (\x\C. A \ B x) = (A \(\x\C. B x))" "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" @@ -1090,7 +1110,7 @@ by auto lemma INT_simps [simp]: - "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \B)" + "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \ B)" "\A B C. (\x\C. A \ B x) = (if C={} then UNIV else A \(\x\C. B x))" "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" @@ -1114,7 +1134,7 @@ lemma UN_extend_simps: "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" - "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" + "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" "\A B C. (A \ (\x\C. B x)) = (\x\C. A \ B x)"