merged
authorberghofe
Tue, 14 Jul 2009 17:18:51 +0200
changeset 32040 830141c9e528
parent 32039 400a519bc888 (current diff)
parent 31996 1d93369079c4 (diff)
child 32041 b09916780820
merged
src/HOL/GCD.thy
--- a/src/HOL/Finite_Set.thy	Tue Jul 14 17:17:37 2009 +0200
+++ b/src/HOL/Finite_Set.thy	Tue Jul 14 17:18:51 2009 +0200
@@ -218,6 +218,12 @@
  "\<lbrakk> finite A; !!M. M \<in> A \<Longrightarrow> finite M \<rbrakk> \<Longrightarrow> finite(\<Union>A)"
 by (induct rule:finite_induct) simp_all
 
+lemma finite_Inter[intro]: "EX A:M. finite(A) \<Longrightarrow> finite(Inter M)"
+by (blast intro: Inter_lower finite_subset)
+
+lemma finite_INT[intro]: "EX x:I. finite(A x) \<Longrightarrow> finite(INT x:I. A x)"
+by (blast intro: INT_lower finite_subset)
+
 lemma finite_empty_induct:
   assumes "finite A"
     and "P A"
@@ -784,6 +790,62 @@
 
 end
 
+context ab_semigroup_idem_mult
+begin
+
+lemma fun_left_comm_idem: "fun_left_comm_idem(op *)"
+apply unfold_locales
+ apply (simp add: mult_ac)
+apply (simp add: mult_idem mult_assoc[symmetric])
+done
+
+end
+
+context lower_semilattice
+begin
+
+lemma ab_semigroup_idem_mult_inf: "ab_semigroup_idem_mult inf"
+proof qed (rule inf_assoc inf_commute inf_idem)+
+
+lemma fold_inf_insert[simp]: "finite A \<Longrightarrow> fold inf b (insert a A) = inf a (fold inf b A)"
+by(rule fun_left_comm_idem.fold_insert_idem[OF ab_semigroup_idem_mult.fun_left_comm_idem[OF ab_semigroup_idem_mult_inf]])
+
+lemma inf_le_fold_inf: "finite A \<Longrightarrow> ALL a:A. b \<le> a \<Longrightarrow> inf b c \<le> fold inf c A"
+by (induct pred:finite) auto
+
+lemma fold_inf_le_inf: "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> fold inf b A \<le> inf a b"
+proof(induct arbitrary: a pred:finite)
+  case empty thus ?case by simp
+next
+  case (insert x A)
+  show ?case
+  proof cases
+    assume "A = {}" thus ?thesis using insert by simp
+  next
+    assume "A \<noteq> {}" thus ?thesis using insert by auto
+  qed
+qed
+
+end
+
+context upper_semilattice
+begin
+
+lemma ab_semigroup_idem_mult_sup: "ab_semigroup_idem_mult sup"
+by (rule lower_semilattice.ab_semigroup_idem_mult_inf)(rule dual_semilattice)
+
+lemma fold_sup_insert[simp]: "finite A \<Longrightarrow> fold sup b (insert a A) = sup a (fold sup b A)"
+by(rule lower_semilattice.fold_inf_insert)(rule dual_semilattice)
+
+lemma fold_sup_le_sup: "finite A \<Longrightarrow> ALL a:A. a \<le> b \<Longrightarrow> fold sup c A \<le> sup b c"
+by(rule lower_semilattice.inf_le_fold_inf)(rule dual_semilattice)
+
+lemma sup_le_fold_sup: "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a b \<le> fold sup b A"
+by(rule lower_semilattice.fold_inf_le_inf)(rule dual_semilattice)
+
+end
+
+
 subsubsection{* The derived combinator @{text fold_image} *}
 
 definition fold_image :: "('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"
@@ -801,7 +863,7 @@
 proof -
   interpret I: fun_left_comm "%x y. (g x) * y"
     by unfold_locales (simp add: mult_ac)
-  show ?thesis using assms by(simp add:fold_image_def I.fold_insert)
+  show ?thesis using assms by(simp add:fold_image_def)
 qed
 
 (*
@@ -829,10 +891,7 @@
 lemma fold_image_reindex:
 assumes fin: "finite A"
 shows "inj_on h A \<Longrightarrow> fold_image times g z (h`A) = fold_image times (g\<circ>h) z A"
-using fin apply induct
- apply simp
-apply simp
-done
+using fin by induct auto
 
 (*
 text{*
@@ -2351,14 +2410,15 @@
 shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
 by(fastsimp simp:surj_def dest!: endo_inj_surj)
 
-corollary infinite_UNIV_nat: "~finite(UNIV::nat set)"
+corollary infinite_UNIV_nat[iff]: "~finite(UNIV::nat set)"
 proof
   assume "finite(UNIV::nat set)"
   with finite_UNIV_inj_surj[of Suc]
   show False by simp (blast dest: Suc_neq_Zero surjD)
 qed
 
-lemma infinite_UNIV_char_0:
+(* Often leads to bogus ATP proofs because of reduced type information, hence noatp *)
+lemma infinite_UNIV_char_0[noatp]:
   "\<not> finite (UNIV::'a::semiring_char_0 set)"
 proof
   assume "finite (UNIV::'a set)"
@@ -2499,13 +2559,6 @@
 context ab_semigroup_idem_mult
 begin
 
-lemma fun_left_comm_idem: "fun_left_comm_idem(op *)"
-apply unfold_locales
- apply (simp add: mult_ac)
-apply (simp add: mult_idem mult_assoc[symmetric])
-done
-
-
 lemma fold1_insert_idem [simp]:
   assumes nonempty: "A \<noteq> {}" and A: "finite A" 
   shows "fold1 times (insert x A) = x * fold1 times A"
@@ -2667,10 +2720,6 @@
 context lower_semilattice
 begin
 
-lemma ab_semigroup_idem_mult_inf:
-  "ab_semigroup_idem_mult inf"
-  proof qed (rule inf_assoc inf_commute inf_idem)+
-
 lemma below_fold1_iff:
   assumes "finite A" "A \<noteq> {}"
   shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
@@ -2716,11 +2765,6 @@
 
 end
 
-lemma (in upper_semilattice) ab_semigroup_idem_mult_sup:
-  "ab_semigroup_idem_mult sup"
-  by (rule lower_semilattice.ab_semigroup_idem_mult_inf)
-    (rule dual_lattice)
-
 context lattice
 begin
 
@@ -2741,7 +2785,7 @@
 apply(erule exE)
 apply(rule order_trans)
 apply(erule (1) fold1_belowI)
-apply(erule (1) lower_semilattice.fold1_belowI [OF dual_lattice])
+apply(erule (1) lower_semilattice.fold1_belowI [OF dual_semilattice])
 done
 
 lemma sup_Inf_absorb [simp]:
@@ -2753,7 +2797,7 @@
 lemma inf_Sup_absorb [simp]:
   "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"
 by (simp add: Sup_fin_def inf_absorb1
-  lower_semilattice.fold1_belowI [OF dual_lattice])
+  lower_semilattice.fold1_belowI [OF dual_semilattice])
 
 end
 
--- a/src/HOL/GCD.thy	Tue Jul 14 17:17:37 2009 +0200
+++ b/src/HOL/GCD.thy	Tue Jul 14 17:18:51 2009 +0200
@@ -37,7 +37,7 @@
 
 subsection {* gcd *}
 
-class gcd = one +
+class gcd = zero + one + dvd +
 
 fixes
   gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" and
@@ -540,15 +540,15 @@
 
 (* to do: add the three variations of these, and for ints? *)
 
-lemma finite_divisors_nat:
-  assumes "(m::nat)~= 0" shows "finite{d. d dvd m}"
+lemma finite_divisors_nat[simp]:
+  assumes "(m::nat) ~= 0" shows "finite{d. d dvd m}"
 proof-
   have "finite{d. d <= m}" by(blast intro: bounded_nat_set_is_finite)
   from finite_subset[OF _ this] show ?thesis using assms
     by(bestsimp intro!:dvd_imp_le)
 qed
 
-lemma finite_divisors_int:
+lemma finite_divisors_int[simp]:
   assumes "(i::int) ~= 0" shows "finite{d. d dvd i}"
 proof-
   have "{d. abs d <= abs i} = {- abs i .. abs i}" by(auto simp:abs_if)
@@ -557,10 +557,25 @@
     by(bestsimp intro!:dvd_imp_le_int)
 qed
 
+lemma Max_divisors_self_nat[simp]: "n\<noteq>0 \<Longrightarrow> Max{d::nat. d dvd n} = n"
+apply(rule antisym)
+ apply (fastsimp intro: Max_le_iff[THEN iffD2] simp: dvd_imp_le)
+apply simp
+done
+
+lemma Max_divisors_self_int[simp]: "n\<noteq>0 \<Longrightarrow> Max{d::int. d dvd n} = abs n"
+apply(rule antisym)
+ apply(rule Max_le_iff[THEN iffD2])
+   apply simp
+  apply fastsimp
+ apply (metis Collect_def abs_ge_self dvd_imp_le_int mem_def zle_trans)
+apply simp
+done
+
 lemma gcd_is_Max_divisors_nat:
   "m ~= 0 \<Longrightarrow> n ~= 0 \<Longrightarrow> gcd (m::nat) n = (Max {d. d dvd m & d dvd n})"
 apply(rule Max_eqI[THEN sym])
-  apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_nat)
+  apply (metis finite_Collect_conjI finite_divisors_nat)
  apply simp
  apply(metis Suc_diff_1 Suc_neq_Zero dvd_imp_le gcd_greatest_iff_nat gcd_pos_nat)
 apply simp
@@ -569,7 +584,7 @@
 lemma gcd_is_Max_divisors_int:
   "m ~= 0 ==> n ~= 0 ==> gcd (m::int) n = (Max {d. d dvd m & d dvd n})"
 apply(rule Max_eqI[THEN sym])
-  apply (metis dvd.eq_iff finite_Collect_conjI finite_divisors_int)
+  apply (metis finite_Collect_conjI finite_divisors_int)
  apply simp
  apply (metis gcd_greatest_iff_int gcd_pos_int zdvd_imp_le)
 apply simp
@@ -1442,31 +1457,61 @@
 lemma lcm_proj1_if_dvd_int [simp]: "(x::int) dvd y \<Longrightarrow> lcm y x = abs y"
 by (subst lcm_commute_int, erule lcm_proj2_if_dvd_int)
 
+lemma lcm_proj1_iff_nat[simp]: "lcm m n = (m::nat) \<longleftrightarrow> n dvd m"
+by (metis lcm_proj1_if_dvd_nat lcm_unique_nat)
+
+lemma lcm_proj2_iff_nat[simp]: "lcm m n = (n::nat) \<longleftrightarrow> m dvd n"
+by (metis lcm_proj2_if_dvd_nat lcm_unique_nat)
+
+lemma lcm_proj1_iff_int[simp]: "lcm m n = abs(m::int) \<longleftrightarrow> n dvd m"
+by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int)
+
+lemma lcm_proj2_iff_int[simp]: "lcm m n = abs(n::int) \<longleftrightarrow> m dvd n"
+by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int)
 
 lemma lcm_assoc_nat: "lcm (lcm n m) (p::nat) = lcm n (lcm m p)"
-apply(rule lcm_unique_nat[THEN iffD1])
-apply (metis dvd.order_trans lcm_unique_nat)
-done
+by(rule lcm_unique_nat[THEN iffD1])(metis dvd.order_trans lcm_unique_nat)
 
 lemma lcm_assoc_int: "lcm (lcm n m) (p::int) = lcm n (lcm m p)"
-apply(rule lcm_unique_int[THEN iffD1])
-apply (metis dvd_trans lcm_unique_int)
-done
+by(rule lcm_unique_int[THEN iffD1])(metis dvd_trans lcm_unique_int)
 
-lemmas lcm_left_commute_nat =
-  mk_left_commute[of lcm, OF lcm_assoc_nat lcm_commute_nat]
-
-lemmas lcm_left_commute_int =
-  mk_left_commute[of lcm, OF lcm_assoc_int lcm_commute_int]
+lemmas lcm_left_commute_nat = mk_left_commute[of lcm, OF lcm_assoc_nat lcm_commute_nat]
+lemmas lcm_left_commute_int = mk_left_commute[of lcm, OF lcm_assoc_int lcm_commute_int]
 
 lemmas lcm_ac_nat = lcm_assoc_nat lcm_commute_nat lcm_left_commute_nat
 lemmas lcm_ac_int = lcm_assoc_int lcm_commute_int lcm_left_commute_int
 
+lemma fun_left_comm_idem_gcd_nat: "fun_left_comm_idem (gcd :: nat\<Rightarrow>nat\<Rightarrow>nat)"
+proof qed (auto simp add: gcd_ac_nat)
+
+lemma fun_left_comm_idem_gcd_int: "fun_left_comm_idem (gcd :: int\<Rightarrow>int\<Rightarrow>int)"
+proof qed (auto simp add: gcd_ac_int)
+
+lemma fun_left_comm_idem_lcm_nat: "fun_left_comm_idem (lcm :: nat\<Rightarrow>nat\<Rightarrow>nat)"
+proof qed (auto simp add: lcm_ac_nat)
+
+lemma fun_left_comm_idem_lcm_int: "fun_left_comm_idem (lcm :: int\<Rightarrow>int\<Rightarrow>int)"
+proof qed (auto simp add: lcm_ac_int)
+
+
+(* FIXME introduce selimattice_bot/top and derive the following lemmas in there: *)
+
+lemma lcm_0_iff_nat[simp]: "lcm (m::nat) n = 0 \<longleftrightarrow> m=0 \<or> n=0"
+by (metis lcm_0_left_nat lcm_0_nat mult_is_0 prod_gcd_lcm_nat)
+
+lemma lcm_0_iff_int[simp]: "lcm (m::int) n = 0 \<longleftrightarrow> m=0 \<or> n=0"
+by (metis lcm_0_int lcm_0_left_int lcm_pos_int zless_le)
+
+lemma lcm_1_iff_nat[simp]: "lcm (m::nat) n = 1 \<longleftrightarrow> m=1 \<and> n=1"
+by (metis gcd_1_nat lcm_unique_nat nat_mult_1 prod_gcd_lcm_nat)
+
+lemma lcm_1_iff_int[simp]: "lcm (m::int) n = 1 \<longleftrightarrow> (m=1 \<or> m = -1) \<and> (n=1 \<or> n = -1)"
+by (auto simp add: abs_mult_self trans [OF lcm_unique_int eq_commute, symmetric] zmult_eq_1_iff)
+
 
 subsection {* Primes *}
 
-(* Is there a better way to handle these, rather than making them
-   elim rules? *)
+(* FIXME Is there a better way to handle these, rather than making them elim rules? *)
 
 lemma prime_ge_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
   by (unfold prime_nat_def, auto)
@@ -1490,7 +1535,7 @@
   by (unfold prime_nat_def, auto)
 
 lemma prime_ge_0_int [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
-  by (unfold prime_int_def prime_nat_def, auto)
+  by (unfold prime_int_def prime_nat_def) auto
 
 lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0"
   by (unfold prime_int_def prime_nat_def, auto)
@@ -1504,8 +1549,6 @@
 lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
   by (unfold prime_int_def prime_nat_def, auto)
 
-thm prime_nat_def;
-thm prime_nat_def [transferred];
 
 lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow>
     m = 1 \<or> m = p))"
@@ -1566,35 +1609,13 @@
 lemma not_prime_eq_prod_nat: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
     EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
   unfolding prime_nat_def dvd_def apply auto
-  apply (subgoal_tac "k > 1")
-  apply force
-  apply (subgoal_tac "k ~= 0")
-  apply force
-  apply (rule notI)
-  apply force
-done
+  by(metis mult_commute linorder_neq_iff linorder_not_le mult_1 n_less_n_mult_m one_le_mult_iff less_imp_le_nat)
 
-(* there's a lot of messing around with signs of products here --
-   could this be made more automatic? *)
 lemma not_prime_eq_prod_int: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
     EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
   unfolding prime_int_altdef dvd_def
   apply auto
-  apply (rule_tac x = m in exI)
-  apply (rule_tac x = k in exI)
-  apply (auto simp add: mult_compare_simps)
-  apply (subgoal_tac "k > 0")
-  apply arith
-  apply (case_tac "k <= 0")
-  apply (subgoal_tac "m * k <= 0")
-  apply force
-  apply (subst zero_compare_simps(8))
-  apply auto
-  apply (subgoal_tac "m * k <= 0")
-  apply force
-  apply (subst zero_compare_simps(8))
-  apply auto
-done
+  by(metis div_mult_self1_is_id div_mult_self2_is_id int_div_less_self int_one_le_iff_zero_less zero_less_mult_pos zless_le)
 
 lemma prime_dvd_power_nat [rule_format]: "prime (p::nat) -->
     n > 0 --> (p dvd x^n --> p dvd x)"
--- a/src/HOL/Lattices.thy	Tue Jul 14 17:17:37 2009 +0200
+++ b/src/HOL/Lattices.thy	Tue Jul 14 17:18:51 2009 +0200
@@ -29,7 +29,7 @@
 
 text {* Dual lattice *}
 
-lemma dual_lattice:
+lemma dual_semilattice:
   "lower_semilattice (op \<ge>) (op >) sup"
 by (rule lower_semilattice.intro, rule dual_order)
   (unfold_locales, simp_all add: sup_least)
@@ -180,6 +180,11 @@
 context lattice
 begin
 
+lemma dual_lattice:
+  "lattice (op \<ge>) (op >) sup inf"
+  by (rule lattice.intro, rule dual_semilattice, rule upper_semilattice.intro, rule dual_order)
+    (unfold_locales, auto)
+
 lemma inf_sup_absorb: "x \<sqinter> (x \<squnion> y) = x"
   by (blast intro: antisym inf_le1 inf_greatest sup_ge1)
 
@@ -252,12 +257,148 @@
  "(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)"
 by(simp add:ACI inf_sup_distrib1)
 
+lemma dual_distrib_lattice:
+  "distrib_lattice (op \<ge>) (op >) sup inf"
+  by (rule distrib_lattice.intro, rule dual_lattice)
+    (unfold_locales, fact inf_sup_distrib1)
+
 lemmas distrib =
   sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2
 
 end
 
 
+subsection {* Boolean algebras *}
+
+class boolean_algebra = distrib_lattice + top + bot + minus + uminus +
+  assumes inf_compl_bot: "x \<sqinter> - x = bot"
+    and sup_compl_top: "x \<squnion> - x = top"
+  assumes diff_eq: "x - y = x \<sqinter> - y"
+begin
+
+lemma dual_boolean_algebra:
+  "boolean_algebra (\<lambda>x y. x \<squnion> - y) uminus (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) top bot"
+  by (rule boolean_algebra.intro, rule dual_distrib_lattice)
+    (unfold_locales,
+      auto simp add: inf_compl_bot sup_compl_top diff_eq less_le_not_le)
+
+lemma compl_inf_bot:
+  "- x \<sqinter> x = bot"
+  by (simp add: inf_commute inf_compl_bot)
+
+lemma compl_sup_top:
+  "- x \<squnion> x = top"
+  by (simp add: sup_commute sup_compl_top)
+
+lemma inf_bot_left [simp]:
+  "bot \<sqinter> x = bot"
+  by (rule inf_absorb1) simp
+
+lemma inf_bot_right [simp]:
+  "x \<sqinter> bot = bot"
+  by (rule inf_absorb2) simp
+
+lemma sup_top_left [simp]:
+  "top \<squnion> x = top"
+  by (rule sup_absorb1) simp
+
+lemma sup_top_right [simp]:
+  "x \<squnion> top = top"
+  by (rule sup_absorb2) simp
+
+lemma inf_top_left [simp]:
+  "top \<sqinter> x = x"
+  by (rule inf_absorb2) simp
+
+lemma inf_top_right [simp]:
+  "x \<sqinter> top = x"
+  by (rule inf_absorb1) simp
+
+lemma sup_bot_left [simp]:
+  "bot \<squnion> x = x"
+  by (rule sup_absorb2) simp
+
+lemma sup_bot_right [simp]:
+  "x \<squnion> bot = x"
+  by (rule sup_absorb1) simp
+
+lemma compl_unique:
+  assumes "x \<sqinter> y = bot"
+    and "x \<squnion> y = top"
+  shows "- x = y"
+proof -
+  have "(x \<sqinter> - x) \<squnion> (- x \<sqinter> y) = (x \<sqinter> y) \<squnion> (- x \<sqinter> y)"
+    using inf_compl_bot assms(1) by simp
+  then have "(- x \<sqinter> x) \<squnion> (- x \<sqinter> y) = (y \<sqinter> x) \<squnion> (y \<sqinter> - x)"
+    by (simp add: inf_commute)
+  then have "- x \<sqinter> (x \<squnion> y) = y \<sqinter> (x \<squnion> - x)"
+    by (simp add: inf_sup_distrib1)
+  then have "- x \<sqinter> top = y \<sqinter> top"
+    using sup_compl_top assms(2) by simp
+  then show "- x = y" by (simp add: inf_top_right)
+qed
+
+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 \<longleftrightarrow> x = y"
+proof
+  assume "- x = - y"
+  then have "- x \<sqinter> y = bot"
+    and "- x \<squnion> y = top"
+    by (simp_all add: compl_inf_bot compl_sup_top)
+  then have "- (- x) = y" by (rule compl_unique)
+  then show "x = y" by simp
+next
+  assume "x = y"
+  then show "- x = - y" by simp
+qed
+
+lemma compl_bot_eq [simp]:
+  "- bot = top"
+proof -
+  from sup_compl_top have "bot \<squnion> - bot = top" .
+  then show ?thesis by simp
+qed
+
+lemma compl_top_eq [simp]:
+  "- top = bot"
+proof -
+  from inf_compl_bot have "top \<sqinter> - top = bot" .
+  then show ?thesis by simp
+qed
+
+lemma compl_inf [simp]:
+  "- (x \<sqinter> y) = - x \<squnion> - y"
+proof (rule compl_unique)
+  have "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = ((x \<sqinter> y) \<sqinter> - x) \<squnion> ((x \<sqinter> y) \<sqinter> - y)"
+    by (rule inf_sup_distrib1)
+  also have "... = (y \<sqinter> (x \<sqinter> - x)) \<squnion> (x \<sqinter> (y \<sqinter> - y))"
+    by (simp only: inf_commute inf_assoc inf_left_commute)
+  finally show "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = bot"
+    by (simp add: inf_compl_bot)
+next
+  have "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = (x \<squnion> (- x \<squnion> - y)) \<sqinter> (y \<squnion> (- x \<squnion> - y))"
+    by (rule sup_inf_distrib2)
+  also have "... = (- y \<squnion> (x \<squnion> - x)) \<sqinter> (- x \<squnion> (y \<squnion> - y))"
+    by (simp only: sup_commute sup_assoc sup_left_commute)
+  finally show "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = top"
+    by (simp add: sup_compl_top)
+qed
+
+lemma compl_sup [simp]:
+  "- (x \<squnion> y) = - x \<sqinter> - y"
+proof -
+  interpret boolean_algebra "\<lambda>x y. x \<squnion> - y" uminus "op \<ge>" "op >" "op \<squnion>" "op \<sqinter>" top bot
+    by (rule dual_boolean_algebra)
+  then show ?thesis by simp
+qed
+
+end
+
+
 subsection {* Uniqueness of inf and sup *}
 
 lemma (in lower_semilattice) inf_unique:
@@ -330,17 +471,24 @@
 
 subsection {* Bool as lattice *}
 
-instantiation bool :: distrib_lattice
+instantiation bool :: boolean_algebra
 begin
 
 definition
+  bool_Compl_def: "uminus = Not"
+
+definition
+  bool_diff_def: "A - B \<longleftrightarrow> A \<and> \<not> B"
+
+definition
   inf_bool_eq: "P \<sqinter> Q \<longleftrightarrow> P \<and> Q"
 
 definition
   sup_bool_eq: "P \<squnion> Q \<longleftrightarrow> P \<or> Q"
 
-instance
-  by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def)
+instance proof
+qed (simp_all add: inf_bool_eq sup_bool_eq le_bool_def
+  bot_bool_eq top_bool_eq bool_Compl_def bool_diff_def, auto)
 
 end
 
@@ -369,7 +517,33 @@
 end
 
 instance "fun" :: (type, distrib_lattice) distrib_lattice
-  by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1)
+proof
+qed (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1)
+
+instantiation "fun" :: (type, uminus) uminus
+begin
+
+definition
+  fun_Compl_def: "- A = (\<lambda>x. - A x)"
+
+instance ..
+
+end
+
+instantiation "fun" :: (type, minus) minus
+begin
+
+definition
+  fun_diff_def: "A - B = (\<lambda>x. A x - B x)"
+
+instance ..
+
+end
+
+instance "fun" :: (type, boolean_algebra) boolean_algebra
+proof
+qed (simp_all add: inf_fun_eq sup_fun_eq bot_fun_eq top_fun_eq fun_Compl_def fun_diff_def
+  inf_compl_bot sup_compl_top diff_eq)
 
 
 text {* redundant bindings *}
--- a/src/HOL/Library/Countable.thy	Tue Jul 14 17:17:37 2009 +0200
+++ b/src/HOL/Library/Countable.thy	Tue Jul 14 17:18:51 2009 +0200
@@ -58,7 +58,7 @@
 subclass (in finite) countable
 proof
   have "finite (UNIV\<Colon>'a set)" by (rule finite_UNIV)
-  with finite_conv_nat_seg_image [of UNIV]
+  with finite_conv_nat_seg_image [of "UNIV::'a set"]
   obtain n and f :: "nat \<Rightarrow> 'a" 
     where "UNIV = f ` {i. i < n}" by auto
   then have "surj f" unfolding surj_def by auto
--- a/src/HOL/Set.thy	Tue Jul 14 17:17:37 2009 +0200
+++ b/src/HOL/Set.thy	Tue Jul 14 17:18:51 2009 +0200
@@ -10,7 +10,6 @@
 
 text {* A set in HOL is simply a predicate. *}
 
-
 subsection {* Basic syntax *}
 
 global
@@ -363,46 +362,6 @@
   Bex_def:      "Bex A P        == EX x. x:A & P(x)"
   Bex1_def:     "Bex1 A P       == EX! x. x:A & P(x)"
 
-instantiation "fun" :: (type, minus) minus
-begin
-
-definition
-  fun_diff_def: "A - B = (%x. A x - B x)"
-
-instance ..
-
-end
-
-instantiation bool :: minus
-begin
-
-definition
-  bool_diff_def: "A - B = (A & ~ B)"
-
-instance ..
-
-end
-
-instantiation "fun" :: (type, uminus) uminus
-begin
-
-definition
-  fun_Compl_def: "- A = (%x. - A x)"
-
-instance ..
-
-end
-
-instantiation bool :: uminus
-begin
-
-definition
-  bool_Compl_def: "- A = (~ A)"
-
-instance ..
-
-end
-
 definition Pow :: "'a set => 'a set set" where
   Pow_def: "Pow A = {B. B \<le> A}"