more concise definition for Inf, Sup on bool
authorhaftmann
Fri, 19 Aug 2011 19:33:31 +0200
changeset 4519343b465f4c480
parent 45164 83c4f8ba0aa3
child 45194 4b5b430eb00e
more concise definition for Inf, Sup on bool
NEWS
src/HOL/Complete_Lattice.thy
     1.1 --- a/NEWS	Fri Aug 19 17:05:10 2011 +0900
     1.2 +++ b/NEWS	Fri Aug 19 19:33:31 2011 +0200
     1.3 @@ -70,7 +70,8 @@
     1.4  generalized theorems INF_cong and SUP_cong.  New type classes for complete
     1.5  boolean algebras and complete linear orders.  Lemmas Inf_less_iff,
     1.6  less_Sup_iff, INF_less_iff, less_SUP_iff now reside in class complete_linorder.
     1.7 -Changed proposition of lemmas Inf_fun_def, Sup_fun_def, Inf_apply, Sup_apply.
     1.8 +Changed proposition of lemmas Inf_bool_def, Sup_bool_def, Inf_fun_def, Sup_fun_def,
     1.9 +Inf_apply, Sup_apply.
    1.10  Redundant lemmas Inf_singleton, Sup_singleton, Inf_binary, Sup_binary,
    1.11  INF_eq, SUP_eq, INF_UNIV_range, SUP_UNIV_range, Int_eq_Inter,
    1.12  INTER_eq_Inter_image, Inter_def, INT_eq, Un_eq_Union, UNION_eq_Union_image,
     2.1 --- a/src/HOL/Complete_Lattice.thy	Fri Aug 19 17:05:10 2011 +0900
     2.2 +++ b/src/HOL/Complete_Lattice.thy	Fri Aug 19 19:33:31 2011 +0200
     2.3 @@ -414,8 +414,7 @@
     2.4    apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf)
     2.5    done
     2.6  
     2.7 -subclass distrib_lattice proof -- {* Question: is it sufficient to include @{class distrib_lattice}
     2.8 -  and prove @{text inf_Sup} and @{text sup_Inf} from that? *}
     2.9 +subclass distrib_lattice proof
    2.10    fix a b c
    2.11    from sup_Inf have "a \<squnion> \<Sqinter>{b, c} = (\<Sqinter>d\<in>{b, c}. a \<squnion> d)" .
    2.12    then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by (simp add: INF_def Inf_insert)
    2.13 @@ -556,13 +555,13 @@
    2.14  begin
    2.15  
    2.16  definition
    2.17 -  "\<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x)"
    2.18 +  [simp]: "\<Sqinter>A \<longleftrightarrow> False \<notin> A"
    2.19  
    2.20  definition
    2.21 -  "\<Squnion>A \<longleftrightarrow> (\<exists>x\<in>A. x)"
    2.22 +  [simp]: "\<Squnion>A \<longleftrightarrow> True \<in> A"
    2.23  
    2.24  instance proof
    2.25 -qed (auto simp add: Inf_bool_def Sup_bool_def)
    2.26 +qed (auto intro: bool_induct)
    2.27  
    2.28  end
    2.29  
    2.30 @@ -572,7 +571,7 @@
    2.31    fix A :: "'a set"
    2.32    fix P :: "'a \<Rightarrow> bool"
    2.33    show "(\<Sqinter>x\<in>A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
    2.34 -    by (auto simp add: Ball_def INF_def Inf_bool_def)
    2.35 +    by (auto simp add: INF_def)
    2.36  qed
    2.37  
    2.38  lemma SUP_bool_eq [simp]:
    2.39 @@ -581,11 +580,11 @@
    2.40    fix A :: "'a set"
    2.41    fix P :: "'a \<Rightarrow> bool"
    2.42    show "(\<Squnion>x\<in>A. P x) \<longleftrightarrow> (\<exists>x\<in>A. P x)"
    2.43 -    by (auto simp add: Bex_def SUP_def Sup_bool_def)
    2.44 +    by (auto simp add: SUP_def)
    2.45  qed
    2.46  
    2.47  instance bool :: complete_boolean_algebra proof
    2.48 -qed (auto simp add: Inf_bool_def Sup_bool_def)
    2.49 +qed (auto intro: bool_induct)
    2.50  
    2.51  instantiation "fun" :: (type, complete_lattice) complete_lattice
    2.52  begin
    2.53 @@ -638,7 +637,7 @@
    2.54    have "(\<forall>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<forall>B\<in>A. x \<in> B)"
    2.55      by auto
    2.56    then show "x \<in> \<Inter>A \<longleftrightarrow> x \<in> {x. \<forall>B \<in> A. x \<in> B}"
    2.57 -    by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def)
    2.58 +    by (simp add: Inf_fun_def) (simp add: mem_def)
    2.59  qed
    2.60  
    2.61  lemma Inter_iff [simp,no_atp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
    2.62 @@ -821,7 +820,7 @@
    2.63    have "(\<exists>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<exists>B\<in>A. x \<in> B)"
    2.64      by auto
    2.65    then show "x \<in> \<Union>A \<longleftrightarrow> x \<in> {x. \<exists>B\<in>A. x \<in> B}"
    2.66 -    by (simp add: Sup_fun_def Sup_bool_def) (simp add: mem_def)
    2.67 +    by (simp add: Sup_fun_def) (simp add: mem_def)
    2.68  qed
    2.69  
    2.70  lemma Union_iff [simp, no_atp]: