--- a/src/HOL/Finite_Set.thy Fri Jul 10 09:24:50 2009 +0200
+++ b/src/HOL/Finite_Set.thy Sun Jul 12 10:14: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_lattice)
+
+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_lattice)
+
+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_lattice)
+
+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_lattice)
+
+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
--- a/src/HOL/GCD.thy Fri Jul 10 09:24:50 2009 +0200
+++ b/src/HOL/GCD.thy Sun Jul 12 10:14:51 2009 +0200
@@ -34,7 +34,7 @@
subsection {* gcd *}
-class gcd = one +
+class gcd = zero + one + dvd +
fixes
gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" and
@@ -537,8 +537,8 @@
(* 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
@@ -1439,31 +1439,46 @@
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)
+
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)
@@ -1487,7 +1502,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)
@@ -1501,8 +1516,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))"
@@ -1563,35 +1576,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/Library/Countable.thy Fri Jul 10 09:24:50 2009 +0200
+++ b/src/HOL/Library/Countable.thy Sun Jul 12 10:14: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