--- a/NEWS Tue Sep 01 14:13:34 2009 +0200
+++ b/NEWS Tue Sep 01 15:39:33 2009 +0200
@@ -18,6 +18,15 @@
*** HOL ***
+* Reorganization of number theory:
+ * former session NumberTheory now names Old_Number_Theory; former session NewNumberTheory
+ named NumberTheory;
+ * split off prime number ingredients from theory GCD to theory Number_Theory/Primes;
+ * moved legacy theories Legacy_GCD and Primes from Library/ to Old_Number_Theory/;
+ * moved theory Pocklington from Library/ to Old_Number_Theory/;
+ * removed various references to Old_Number_Theory from HOL distribution.
+INCOMPATIBILITY.
+
* New testing tool "Mirabelle" for automated (proof) tools. Applies
several tools and tactics like sledgehammer, metis, or quickcheck, to
every proof step in a theory. To be used in batch mode via the
--- a/src/HOL/Extraction/Euclid.thy Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/Extraction/Euclid.thy Tue Sep 01 15:39:33 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Extraction/Euclid.thy
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
Freek Wiedijk, Radboud University Nijmegen
Stefan Berghofer, TU Muenchen
@@ -8,7 +7,7 @@
header {* Euclid's theorem *}
theory Euclid
-imports "~~/src/HOL/NumberTheory/Factorization" Util Efficient_Nat
+imports "~~/src/HOL/Old_Number_Theory/Factorization" Util Efficient_Nat
begin
text {*
--- a/src/HOL/Extraction/ROOT.ML Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/Extraction/ROOT.ML Tue Sep 01 15:39:33 2009 +0200
@@ -1,5 +1,4 @@
(* Title: HOL/Extraction/ROOT.ML
- ID: $Id$
Examples for program extraction in Higher-Order Logic.
*)
@@ -8,5 +7,5 @@
warning "HOL proof terms required for running extraction examples"
else
(Proofterm.proofs := 2;
- no_document use_thys ["Efficient_Nat", "~~/src/HOL/NumberTheory/Factorization"];
+ no_document use_thys ["Efficient_Nat", "~~/src/HOL/Old_Number_Theory/Factorization"];
use_thys ["Greatest_Common_Divisor", "Warshall", "Higman", "Pigeonhole", "Euclid"]);
--- a/src/HOL/GCD.thy Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/GCD.thy Tue Sep 01 15:39:33 2009 +0200
@@ -1,11 +1,9 @@
-(* Title: GCD.thy
- Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+(* Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow
-This file deals with the functions gcd and lcm, and properties of
-primes. Definitions and lemmas are proved uniformly for the natural
-numbers and integers.
+This file deals with the functions gcd and lcm. Definitions and
+lemmas are proved uniformly for the natural numbers and integers.
This file combines and revises a number of prior developments.
@@ -52,11 +50,6 @@
end
-class prime = one +
-
-fixes
- prime :: "'a \<Rightarrow> bool"
-
(* definitions for the natural numbers *)
@@ -80,20 +73,6 @@
end
-instantiation nat :: prime
-
-begin
-
-definition
- prime_nat :: "nat \<Rightarrow> bool"
-where
- [code del]: "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-
-instance proof qed
-
-end
-
-
(* definitions for the integers *)
instantiation int :: gcd
@@ -115,28 +94,13 @@
end
-instantiation int :: prime
-
-begin
-
-definition
- prime_int :: "int \<Rightarrow> bool"
-where
- [code del]: "prime_int p = prime (nat p)"
-
-instance proof qed
-
-end
-
-
subsection {* Set up Transfer *}
lemma transfer_nat_int_gcd:
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> gcd (nat x) (nat y) = nat (gcd x y)"
"(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> lcm (nat x) (nat y) = nat (lcm x y)"
- "(x::int) >= 0 \<Longrightarrow> prime (nat x) = prime x"
- unfolding gcd_int_def lcm_int_def prime_int_def
+ unfolding gcd_int_def lcm_int_def
by auto
lemma transfer_nat_int_gcd_closures:
@@ -150,8 +114,7 @@
lemma transfer_int_nat_gcd:
"gcd (int x) (int y) = int (gcd x y)"
"lcm (int x) (int y) = int (lcm x y)"
- "prime (int x) = prime x"
- by (unfold gcd_int_def lcm_int_def prime_int_def, auto)
+ by (unfold gcd_int_def lcm_int_def, auto)
lemma transfer_int_nat_gcd_closures:
"is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> gcd x y >= 0"
@@ -1003,20 +966,6 @@
apply (auto simp add: gcd_mult_cancel_int)
done
-lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
- unfolding prime_nat_def
- apply (subst even_mult_two_ex)
- apply clarify
- apply (drule_tac x = 2 in spec)
- apply auto
-done
-
-lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
- unfolding prime_int_def
- apply (frule prime_odd_nat)
- apply (auto simp add: even_nat_def)
-done
-
lemma coprime_common_divisor_nat: "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow>
x dvd b \<Longrightarrow> x = 1"
apply (subgoal_tac "x dvd gcd a b")
@@ -1753,327 +1702,4 @@
show ?thesis by(simp add: Gcd_def fold_set gcd_commute_int)
qed
-
-subsection {* Primes *}
-
-(* 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)
-
-lemma prime_gt_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
- by (unfold prime_nat_def, auto)
-
-lemma prime_ge_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
- by (unfold prime_nat_def, auto)
-
-lemma prime_gt_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
- by (unfold prime_nat_def, auto)
-
-lemma prime_ge_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
- by (unfold prime_nat_def, auto)
-
-lemma prime_gt_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
- by (unfold prime_nat_def, auto)
-
-lemma prime_ge_2_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
- 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
-
-lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0"
- by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_ge_1_int [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
- by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_gt_1_int [elim]: "prime (p::int) \<Longrightarrow> p > 1"
- by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
- by (unfold prime_int_def prime_nat_def, auto)
-
-
-lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow>
- m = 1 \<or> m = p))"
- using prime_nat_def [transferred]
- apply (case_tac "p >= 0")
- by (blast, auto simp add: prime_ge_0_int)
-
-lemma prime_imp_coprime_nat: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
- apply (unfold prime_nat_def)
- apply (metis gcd_dvd1_nat gcd_dvd2_nat)
- done
-
-lemma prime_imp_coprime_int: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
- apply (unfold prime_int_altdef)
- apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int)
- done
-
-lemma prime_dvd_mult_nat: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
- by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat)
-
-lemma prime_dvd_mult_int: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
- by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int)
-
-lemma prime_dvd_mult_eq_nat [simp]: "prime (p::nat) \<Longrightarrow>
- p dvd m * n = (p dvd m \<or> p dvd n)"
- by (rule iffI, rule prime_dvd_mult_nat, auto)
-
-lemma prime_dvd_mult_eq_int [simp]: "prime (p::int) \<Longrightarrow>
- p dvd m * n = (p dvd m \<or> p dvd n)"
- by (rule iffI, rule prime_dvd_mult_int, auto)
-
-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
- 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)
-
-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
- 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)"
- by (induct n rule: nat_induct, auto)
-
-lemma prime_dvd_power_int [rule_format]: "prime (p::int) -->
- n > 0 --> (p dvd x^n --> p dvd x)"
- apply (induct n rule: nat_induct, auto)
- apply (frule prime_ge_0_int)
- apply auto
-done
-
-subsubsection{* Make prime naively executable *}
-
-lemma zero_not_prime_nat [simp]: "~prime (0::nat)"
- by (simp add: prime_nat_def)
-
-lemma zero_not_prime_int [simp]: "~prime (0::int)"
- by (simp add: prime_int_def)
-
-lemma one_not_prime_nat [simp]: "~prime (1::nat)"
- by (simp add: prime_nat_def)
-
-lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
- by (simp add: prime_nat_def One_nat_def)
-
-lemma one_not_prime_int [simp]: "~prime (1::int)"
- by (simp add: prime_int_def)
-
-lemma prime_nat_code[code]:
- "prime(p::nat) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))"
-apply(simp add: Ball_def)
-apply (metis less_not_refl prime_nat_def dvd_triv_right not_prime_eq_prod_nat)
-done
-
-lemma prime_nat_simp:
- "prime(p::nat) = (p > 1 & (list_all (%n. ~ n dvd p) [2..<p]))"
-apply(simp only:prime_nat_code list_ball_code greaterThanLessThan_upt)
-apply(simp add:nat_number One_nat_def)
-done
-
-lemmas prime_nat_simp_number_of[simp] = prime_nat_simp[of "number_of m", standard]
-
-lemma prime_int_code[code]:
- "prime(p::int) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))" (is "?L = ?R")
-proof
- assume "?L" thus "?R"
- by (clarsimp simp: prime_gt_1_int) (metis int_one_le_iff_zero_less prime_int_altdef zless_le)
-next
- assume "?R" thus "?L" by (clarsimp simp:Ball_def) (metis dvdI not_prime_eq_prod_int)
-qed
-
-lemma prime_int_simp:
- "prime(p::int) = (p > 1 & (list_all (%n. ~ n dvd p) [2..p - 1]))"
-apply(simp only:prime_int_code list_ball_code greaterThanLessThan_upto)
-apply simp
-done
-
-lemmas prime_int_simp_number_of[simp] = prime_int_simp[of "number_of m", standard]
-
-lemma two_is_prime_nat [simp]: "prime (2::nat)"
-by simp
-
-lemma two_is_prime_int [simp]: "prime (2::int)"
-by simp
-
-text{* A bit of regression testing: *}
-
-lemma "prime(97::nat)"
-by simp
-
-lemma "prime(97::int)"
-by simp
-
-lemma "prime(997::nat)"
-by eval
-
-lemma "prime(997::int)"
-by eval
-
-
-lemma prime_imp_power_coprime_nat: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
- apply (rule coprime_exp_nat)
- apply (subst gcd_commute_nat)
- apply (erule (1) prime_imp_coprime_nat)
-done
-
-lemma prime_imp_power_coprime_int: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
- apply (rule coprime_exp_int)
- apply (subst gcd_commute_int)
- apply (erule (1) prime_imp_coprime_int)
-done
-
-lemma primes_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
- apply (rule prime_imp_coprime_nat, assumption)
- apply (unfold prime_nat_def, auto)
-done
-
-lemma primes_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
- apply (rule prime_imp_coprime_int, assumption)
- apply (unfold prime_int_altdef, clarify)
- apply (drule_tac x = q in spec)
- apply (drule_tac x = p in spec)
- apply auto
-done
-
-lemma primes_imp_powers_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
- by (rule coprime_exp2_nat, rule primes_coprime_nat)
-
-lemma primes_imp_powers_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
- by (rule coprime_exp2_int, rule primes_coprime_int)
-
-lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
- apply (induct n rule: nat_less_induct)
- apply (case_tac "n = 0")
- using two_is_prime_nat apply blast
- apply (case_tac "prime n")
- apply blast
- apply (subgoal_tac "n > 1")
- apply (frule (1) not_prime_eq_prod_nat)
- apply (auto intro: dvd_mult dvd_mult2)
-done
-
-(* An Isar version:
-
-lemma prime_factor_b_nat:
- fixes n :: nat
- assumes "n \<noteq> 1"
- shows "\<exists>p. prime p \<and> p dvd n"
-
-using `n ~= 1`
-proof (induct n rule: less_induct_nat)
- fix n :: nat
- assume "n ~= 1" and
- ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)"
- thus "\<exists>p. prime p \<and> p dvd n"
- proof -
- {
- assume "n = 0"
- moreover note two_is_prime_nat
- ultimately have ?thesis
- by (auto simp del: two_is_prime_nat)
- }
- moreover
- {
- assume "prime n"
- hence ?thesis by auto
- }
- moreover
- {
- assume "n ~= 0" and "~ prime n"
- with `n ~= 1` have "n > 1" by auto
- with `~ prime n` and not_prime_eq_prod_nat obtain m k where
- "n = m * k" and "1 < m" and "m < n" by blast
- with ih obtain p where "prime p" and "p dvd m" by blast
- with `n = m * k` have ?thesis by auto
- }
- ultimately show ?thesis by blast
- qed
-qed
-
-*)
-
-text {* One property of coprimality is easier to prove via prime factors. *}
-
-lemma prime_divprod_pow_nat:
- assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b"
- shows "p^n dvd a \<or> p^n dvd b"
-proof-
- {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
- apply (cases "n=0", simp_all)
- apply (cases "a=1", simp_all) done}
- moreover
- {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
- then obtain m where m: "n = Suc m" by (cases n, auto)
- from n have "p dvd p^n" by (intro dvd_power, auto)
- also note pab
- finally have pab': "p dvd a * b".
- from prime_dvd_mult_nat[OF p pab']
- have "p dvd a \<or> p dvd b" .
- moreover
- {assume pa: "p dvd a"
- have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> p dvd b" by auto
- with p have "coprime b p"
- by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
- hence pnb: "coprime (p^n) b"
- by (subst gcd_commute_nat, rule coprime_exp_nat)
- from coprime_divprod_nat[OF pnba pnb] have ?thesis by blast }
- moreover
- {assume pb: "p dvd b"
- have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from coprime_common_divisor_nat [OF ab, of p] pb p have "\<not> p dvd a"
- by auto
- with p have "coprime a p"
- by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
- hence pna: "coprime (p^n) a"
- by (subst gcd_commute_nat, rule coprime_exp_nat)
- from coprime_divprod_nat[OF pab pna] have ?thesis by blast }
- ultimately have ?thesis by blast}
- ultimately show ?thesis by blast
-qed
-
-subsection {* Infinitely many primes *}
-
-lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1"
-proof-
- have f1: "fact n + 1 \<noteq> 1" using fact_ge_one_nat [of n] by arith
- from prime_factor_nat [OF f1]
- obtain p where "prime p" and "p dvd fact n + 1" by auto
- hence "p \<le> fact n + 1"
- by (intro dvd_imp_le, auto)
- {assume "p \<le> n"
- from `prime p` have "p \<ge> 1"
- by (cases p, simp_all)
- with `p <= n` have "p dvd fact n"
- by (intro dvd_fact_nat)
- with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n"
- by (rule dvd_diff_nat)
- hence "p dvd 1" by simp
- hence "p <= 1" by auto
- moreover from `prime p` have "p > 1" by auto
- ultimately have False by auto}
- hence "n < p" by arith
- with `prime p` and `p <= fact n + 1` show ?thesis by auto
-qed
-
-lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
-using next_prime_bound by auto
-
-lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
-proof
- assume "finite {(p::nat). prime p}"
- with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
- by auto
- then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
- by auto
- with bigger_prime [of b] show False by auto
-qed
-
-
end
--- a/src/HOL/Import/HOL4Compat.thy Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/Import/HOL4Compat.thy Tue Sep 01 15:39:33 2009 +0200
@@ -3,7 +3,7 @@
*)
theory HOL4Compat
-imports HOL4Setup Complex_Main Primes ContNotDenum
+imports HOL4Setup Complex_Main "~~/src/HOL/Old_Number_Theory/Primes" ContNotDenum
begin
no_notation differentiable (infixl "differentiable" 60)
--- a/src/HOL/IsaMakefile Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/IsaMakefile Tue Sep 01 15:39:33 2009 +0200
@@ -34,8 +34,8 @@
HOL-Modelcheck \
HOL-NanoJava \
HOL-Nominal-Examples \
- HOL-NewNumberTheory \
- HOL-NumberTheory \
+ HOL-Number_Theory \
+ HOL-Old_Number_Theory \
HOL-Prolog \
HOL-SET-Protocol \
HOL-SizeChange \
@@ -282,11 +282,13 @@
Complex.thy \
Deriv.thy \
Fact.thy \
+ GCD.thy \
Integration.thy \
Lim.thy \
Limits.thy \
Ln.thy \
Log.thy \
+ Lubs.thy \
MacLaurin.thy \
NatTransfer.thy \
NthRoot.thy \
@@ -294,9 +296,7 @@
Series.thy \
Taylor.thy \
Transcendental.thy \
- GCD.thy \
Parity.thy \
- Lubs.thy \
PReal.thy \
Rational.thy \
RComplete.thy \
@@ -330,10 +330,10 @@
Library/Finite_Cartesian_Product.thy Library/FrechetDeriv.thy \
Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy \
Library/Inner_Product.thy Library/Kleene_Algebra.thy \
- Library/Lattice_Syntax.thy Library/Legacy_GCD.thy \
+ Library/Lattice_Syntax.thy \
Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy \
Library/State_Monad.thy Library/Nat_Int_Bij.thy Library/Multiset.thy \
- Library/Permutation.thy Library/Primes.thy Library/Pocklington.thy \
+ Library/Permutation.thy \
Library/Quotient.thy Library/Quicksort.thy Library/Nat_Infinity.thy \
Library/Word.thy Library/README.html Library/Continuity.thy \
Library/Order_Relation.thy Library/Nested_Environment.thy \
@@ -487,38 +487,39 @@
@cd Import/HOLLight; $(ISABELLE_TOOL) usedir -b $(OUT)/HOL HOLLight
-## HOL-NewNumberTheory
+## HOL-Number_Theory
-HOL-NewNumberTheory: HOL $(LOG)/HOL-NewNumberTheory.gz
+HOL-Number_Theory: HOL $(LOG)/HOL-Number_Theory.gz
-$(LOG)/HOL-NewNumberTheory.gz: $(OUT)/HOL $(ALGEBRA_DEPENDENCIES) \
+$(LOG)/HOL-Number_Theory.gz: $(OUT)/HOL $(ALGEBRA_DEPENDENCIES) \
Library/Multiset.thy \
- NewNumberTheory/Binomial.thy \
- NewNumberTheory/Cong.thy \
- NewNumberTheory/Fib.thy \
- NewNumberTheory/MiscAlgebra.thy \
- NewNumberTheory/Residues.thy \
- NewNumberTheory/UniqueFactorization.thy \
- NewNumberTheory/ROOT.ML
- @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL NewNumberTheory
+ Number_Theory/Binomial.thy \
+ Number_Theory/Cong.thy \
+ Number_Theory/Fib.thy \
+ Number_Theory/MiscAlgebra.thy \
+ Number_Theory/Number_Theory.thy \
+ Number_Theory/Residues.thy \
+ Number_Theory/UniqueFactorization.thy \
+ Number_Theory/ROOT.ML
+ @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Number_Theory
-## HOL-NumberTheory
+## HOL-Old_Number_Theory
-HOL-NumberTheory: HOL $(LOG)/HOL-NumberTheory.gz
+HOL-Old_Number_Theory: HOL $(LOG)/HOL-Old_Number_Theory.gz
-$(LOG)/HOL-NumberTheory.gz: $(OUT)/HOL Library/Permutation.thy \
- Library/Primes.thy NumberTheory/Fib.thy \
- NumberTheory/Factorization.thy NumberTheory/BijectionRel.thy \
- NumberTheory/Chinese.thy NumberTheory/EulerFermat.thy \
- NumberTheory/IntFact.thy NumberTheory/IntPrimes.thy \
- NumberTheory/WilsonBij.thy NumberTheory/WilsonRuss.thy \
- NumberTheory/Finite2.thy NumberTheory/Int2.thy \
- NumberTheory/EvenOdd.thy NumberTheory/Residues.thy \
- NumberTheory/Euler.thy NumberTheory/Gauss.thy \
- NumberTheory/Quadratic_Reciprocity.thy Library/Infinite_Set.thy \
- NumberTheory/ROOT.ML
- @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL NumberTheory
+$(LOG)/HOL-Old_Number_Theory.gz: $(OUT)/HOL Library/Permutation.thy \
+ Old_Number_Theory/Primes.thy Old_Number_Theory/Fib.thy \
+ Old_Number_Theory/Factorization.thy Old_Number_Theory/BijectionRel.thy \
+ Old_Number_Theory/Chinese.thy Old_Number_Theory/EulerFermat.thy \
+ Old_Number_Theory/IntFact.thy Old_Number_Theory/IntPrimes.thy \
+ Old_Number_Theory/WilsonBij.thy Old_Number_Theory/WilsonRuss.thy \
+ Old_Number_Theory/Finite2.thy Old_Number_Theory/Int2.thy \
+ Old_Number_Theory/EvenOdd.thy Old_Number_Theory/Residues.thy \
+ Old_Number_Theory/Euler.thy Old_Number_Theory/Gauss.thy \
+ Old_Number_Theory/Quadratic_Reciprocity.thy Library/Infinite_Set.thy \
+ Old_Number_Theory/Legacy_GCD.thy Old_Number_Theory/Pocklington.thy Old_Number_Theory/ROOT.ML
+ @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Old_Number_Theory
## HOL-Hoare
@@ -573,7 +574,7 @@
Library/FuncSet.thy \
Library/Multiset.thy \
Library/Permutation.thy \
- Library/Primes.thy \
+ Number_Theory/Primes.thy \
Algebra/AbelCoset.thy \
Algebra/Bij.thy \
Algebra/Congruence.thy \
@@ -876,7 +877,7 @@
HOL-ex: HOL $(LOG)/HOL-ex.gz
$(LOG)/HOL-ex.gz: $(OUT)/HOL Library/Commutative_Ring.thy \
- Library/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy \
+ Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy \
ex/Arith_Examples.thy \
ex/Arithmetic_Series_Complex.thy ex/BT.thy ex/BinEx.thy \
ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy \
--- a/src/HOL/Isar_examples/ROOT.ML Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/Isar_examples/ROOT.ML Tue Sep 01 15:39:33 2009 +0200
@@ -4,7 +4,7 @@
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
*)
-no_document use_thys ["../NumberTheory/Primes", "../NumberTheory/Fibonacci"];
+no_document use_thys ["../Old_Number_Theory/Primes", "../Old_Number_Theory/Fibonacci"];
use_thys [
"Basic_Logic",
--- a/src/HOL/Library/Legacy_GCD.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,787 +0,0 @@
-(* Title: HOL/GCD.thy
- Author: Christophe Tabacznyj and Lawrence C Paulson
- Copyright 1996 University of Cambridge
-*)
-
-header {* The Greatest Common Divisor *}
-
-theory Legacy_GCD
-imports Main
-begin
-
-text {*
- See \cite{davenport92}. \bigskip
-*}
-
-subsection {* Specification of GCD on nats *}
-
-definition
- is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
- [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
- (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
-
-text {* Uniqueness *}
-
-lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
- by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
-
-text {* Connection to divides relation *}
-
-lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
- by (auto simp add: is_gcd_def)
-
-text {* Commutativity *}
-
-lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
- by (auto simp add: is_gcd_def)
-
-
-subsection {* GCD on nat by Euclid's algorithm *}
-
-fun
- gcd :: "nat => nat => nat"
-where
- "gcd m n = (if n = 0 then m else gcd n (m mod n))"
-lemma gcd_induct [case_names "0" rec]:
- fixes m n :: nat
- assumes "\<And>m. P m 0"
- and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
- shows "P m n"
-proof (induct m n rule: gcd.induct)
- case (1 m n) with assms show ?case by (cases "n = 0") simp_all
-qed
-
-lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
- by simp
-
-lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
- by simp
-
-lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
- by simp
-
-lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
- by simp
-
-lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
- unfolding One_nat_def by (rule gcd_1)
-
-declare gcd.simps [simp del]
-
-text {*
- \medskip @{term "gcd m n"} divides @{text m} and @{text n}. The
- conjunctions don't seem provable separately.
-*}
-
-lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
- and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
- apply (induct m n rule: gcd_induct)
- apply (simp_all add: gcd_non_0)
- apply (blast dest: dvd_mod_imp_dvd)
- done
-
-text {*
- \medskip Maximality: for all @{term m}, @{term n}, @{term k}
- naturals, if @{term k} divides @{term m} and @{term k} divides
- @{term n} then @{term k} divides @{term "gcd m n"}.
-*}
-
-lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
- by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
-
-text {*
- \medskip Function gcd yields the Greatest Common Divisor.
-*}
-
-lemma is_gcd: "is_gcd m n (gcd m n) "
- by (simp add: is_gcd_def gcd_greatest)
-
-
-subsection {* Derived laws for GCD *}
-
-lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
- by (blast intro!: gcd_greatest intro: dvd_trans)
-
-lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
- by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
-
-lemma gcd_commute: "gcd m n = gcd n m"
- apply (rule is_gcd_unique)
- apply (rule is_gcd)
- apply (subst is_gcd_commute)
- apply (simp add: is_gcd)
- done
-
-lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
- apply (rule is_gcd_unique)
- apply (rule is_gcd)
- apply (simp add: is_gcd_def)
- apply (blast intro: dvd_trans)
- done
-
-lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
- by (simp add: gcd_commute)
-
-lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
- unfolding One_nat_def by (rule gcd_1_left)
-
-text {*
- \medskip Multiplication laws
-*}
-
-lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
- -- {* \cite[page 27]{davenport92} *}
- apply (induct m n rule: gcd_induct)
- apply simp
- apply (case_tac "k = 0")
- apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
- done
-
-lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
- apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
- done
-
-lemma gcd_self [simp, algebra]: "gcd k k = k"
- apply (rule gcd_mult [of k 1, simplified])
- done
-
-lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
- apply (insert gcd_mult_distrib2 [of m k n])
- apply simp
- apply (erule_tac t = m in ssubst)
- apply simp
- done
-
-lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
- by (auto intro: relprime_dvd_mult dvd_mult2)
-
-lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
- apply (rule dvd_anti_sym)
- apply (rule gcd_greatest)
- apply (rule_tac n = k in relprime_dvd_mult)
- apply (simp add: gcd_assoc)
- apply (simp add: gcd_commute)
- apply (simp_all add: mult_commute)
- done
-
-
-text {* \medskip Addition laws *}
-
-lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
- by (cases "n = 0") (auto simp add: gcd_non_0)
-
-lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
-proof -
- have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
- also have "... = gcd (n + m) m" by (simp add: add_commute)
- also have "... = gcd n m" by simp
- also have "... = gcd m n" by (rule gcd_commute)
- finally show ?thesis .
-qed
-
-lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
- apply (subst add_commute)
- apply (rule gcd_add2)
- done
-
-lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
- by (induct k) (simp_all add: add_assoc)
-
-lemma gcd_dvd_prod: "gcd m n dvd m * n"
- using mult_dvd_mono [of 1] by auto
-
-text {*
- \medskip Division by gcd yields rrelatively primes.
-*}
-
-lemma div_gcd_relprime:
- assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
- shows "gcd (a div gcd a b) (b div gcd a b) = 1"
-proof -
- let ?g = "gcd a b"
- let ?a' = "a div ?g"
- let ?b' = "b div ?g"
- let ?g' = "gcd ?a' ?b'"
- have dvdg: "?g dvd a" "?g dvd b" by simp_all
- have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
- from dvdg dvdg' obtain ka kb ka' kb' where
- kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
- unfolding dvd_def by blast
- then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
- then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
- by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
- dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
- have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
- then have gp: "?g > 0" by simp
- from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
- with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
-qed
-
-
-lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
-proof(auto)
- assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
- from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b]
- have th: "gcd a b dvd d" by blast
- from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast
-qed
-
-lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
- shows "gcd x y = gcd u v"
-proof-
- from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
- with gcd_unique[of "gcd u v" x y] show ?thesis by auto
-qed
-
-lemma ind_euclid:
- assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0"
- and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)"
- shows "P a b"
-proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
- fix n a b
- assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
- have "a = b \<or> a < b \<or> b < a" by arith
- moreover {assume eq: "a= b"
- from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
- moreover
- {assume lt: "a < b"
- hence "a + b - a < n \<or> a = 0" using H(2) by arith
- moreover
- {assume "a =0" with z c have "P a b" by blast }
- moreover
- {assume ab: "a + b - a < n"
- have th0: "a + b - a = a + (b - a)" using lt by arith
- from add[rule_format, OF H(1)[rule_format, OF ab th0]]
- have "P a b" by (simp add: th0[symmetric])}
- ultimately have "P a b" by blast}
- moreover
- {assume lt: "a > b"
- hence "b + a - b < n \<or> b = 0" using H(2) by arith
- moreover
- {assume "b =0" with z c have "P a b" by blast }
- moreover
- {assume ab: "b + a - b < n"
- have th0: "b + a - b = b + (a - b)" using lt by arith
- from add[rule_format, OF H(1)[rule_format, OF ab th0]]
- have "P b a" by (simp add: th0[symmetric])
- hence "P a b" using c by blast }
- ultimately have "P a b" by blast}
-ultimately show "P a b" by blast
-qed
-
-lemma bezout_lemma:
- assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
- shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
-using ex
-apply clarsimp
-apply (rule_tac x="d" in exI, simp add: dvd_add)
-apply (case_tac "a * x = b * y + d" , simp_all)
-apply (rule_tac x="x + y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x + y" in exI)
-apply algebra
-done
-
-lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
-apply(induct a b rule: ind_euclid)
-apply blast
-apply clarify
-apply (rule_tac x="a" in exI, simp add: dvd_add)
-apply clarsimp
-apply (rule_tac x="d" in exI)
-apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
-apply (rule_tac x="x+y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x+y" in exI)
-apply algebra
-done
-
-lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
-using bezout_add[of a b]
-apply clarsimp
-apply (rule_tac x="d" in exI, simp)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="y" in exI)
-apply auto
-done
-
-
-text {* We can get a stronger version with a nonzeroness assumption. *}
-lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
-
-lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
- shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
-proof-
- from nz have ap: "a > 0" by simp
- from bezout_add[of a b]
- have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
- moreover
- {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
- from H have ?thesis by blast }
- moreover
- {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
- {assume b0: "b = 0" with H have ?thesis by simp}
- moreover
- {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
- from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
- moreover
- {assume db: "d=b"
- from prems have ?thesis apply simp
- apply (rule exI[where x = b], simp)
- apply (rule exI[where x = b])
- by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
- moreover
- {assume db: "d < b"
- {assume "x=0" hence ?thesis using prems by simp }
- moreover
- {assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
-
- from db have "d \<le> b - 1" by simp
- hence "d*b \<le> b*(b - 1)" by simp
- with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
- have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
- from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
- hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
- hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
- by (simp only: diff_add_assoc[OF dble, of d, symmetric])
- hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
- by (simp only: diff_mult_distrib2 add_commute mult_ac)
- hence ?thesis using H(1,2)
- apply -
- apply (rule exI[where x=d], simp)
- apply (rule exI[where x="(b - 1) * y"])
- by (rule exI[where x="x*(b - 1) - d"], simp)}
- ultimately have ?thesis by blast}
- ultimately have ?thesis by blast}
- ultimately have ?thesis by blast}
- ultimately show ?thesis by blast
-qed
-
-
-lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
-proof-
- let ?g = "gcd a b"
- from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
- from d(1,2) have "d dvd ?g" by simp
- then obtain k where k: "?g = d*k" unfolding dvd_def by blast
- from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast
- hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k"
- by (algebra add: diff_mult_distrib)
- hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g"
- by (simp add: k mult_assoc)
- thus ?thesis by blast
-qed
-
-lemma bezout_gcd_strong: assumes a: "a \<noteq> 0"
- shows "\<exists>x y. a * x = b * y + gcd a b"
-proof-
- let ?g = "gcd a b"
- from bezout_add_strong[OF a, of b]
- obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
- from d(1,2) have "d dvd ?g" by simp
- then obtain k where k: "?g = d*k" unfolding dvd_def by blast
- from d(3) have "a * x * k = (b * y + d) *k " by algebra
- hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
- thus ?thesis by blast
-qed
-
-lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
-by(simp add: gcd_mult_distrib2 mult_commute)
-
-lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
- (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
- let ?g = "gcd a b"
- {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
- from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
- by blast
- hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
- hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k"
- by (simp only: diff_mult_distrib)
- hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
- by (simp add: k[symmetric] mult_assoc)
- hence ?lhs by blast}
- moreover
- {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
- have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
- using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
- from dvd_diff_nat[OF dv(1,2)] dvd_diff_nat[OF dv(3,4)] H
- have ?rhs by auto}
- ultimately show ?thesis by blast
-qed
-
-lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
-proof-
- let ?g = "gcd a b"
- have dv: "?g dvd a*x" "?g dvd b * y"
- using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
- from dvd_add[OF dv] H
- show ?thesis by auto
-qed
-
-lemma gcd_mult': "gcd b (a * b) = b"
-by (simp add: gcd_mult mult_commute[of a b])
-
-lemma gcd_add: "gcd(a + b) b = gcd a b"
- "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
-apply (simp_all add: gcd_add1)
-by (simp add: gcd_commute gcd_add1)
-
-lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
-proof-
- {fix a b assume H: "b \<le> (a::nat)"
- hence th: "a - b + b = a" by arith
- from gcd_add(1)[of "a - b" b] th have "gcd(a - b) b = gcd a b" by simp}
- note th = this
-{
- assume ab: "b \<le> a"
- from th[OF ab] show "gcd (a - b) b = gcd a b" by blast
-next
- assume ab: "a \<le> b"
- from th[OF ab] show "gcd a (b - a) = gcd a b"
- by (simp add: gcd_commute)}
-qed
-
-
-subsection {* LCM defined by GCD *}
-
-
-definition
- lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
- lcm_def: "lcm m n = m * n div gcd m n"
-
-lemma prod_gcd_lcm:
- "m * n = gcd m n * lcm m n"
- unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
-
-lemma lcm_0 [simp]: "lcm m 0 = 0"
- unfolding lcm_def by simp
-
-lemma lcm_1 [simp]: "lcm m 1 = m"
- unfolding lcm_def by simp
-
-lemma lcm_0_left [simp]: "lcm 0 n = 0"
- unfolding lcm_def by simp
-
-lemma lcm_1_left [simp]: "lcm 1 m = m"
- unfolding lcm_def by simp
-
-lemma dvd_pos:
- fixes n m :: nat
- assumes "n > 0" and "m dvd n"
- shows "m > 0"
-using assms by (cases m) auto
-
-lemma lcm_least:
- assumes "m dvd k" and "n dvd k"
- shows "lcm m n dvd k"
-proof (cases k)
- case 0 then show ?thesis by auto
-next
- case (Suc _) then have pos_k: "k > 0" by auto
- from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
- with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
- from assms obtain p where k_m: "k = m * p" using dvd_def by blast
- from assms obtain q where k_n: "k = n * q" using dvd_def by blast
- from pos_k k_m have pos_p: "p > 0" by auto
- from pos_k k_n have pos_q: "q > 0" by auto
- have "k * k * gcd q p = k * gcd (k * q) (k * p)"
- by (simp add: mult_ac gcd_mult_distrib2)
- also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
- by (simp add: k_m [symmetric] k_n [symmetric])
- also have "\<dots> = k * p * q * gcd m n"
- by (simp add: mult_ac gcd_mult_distrib2)
- finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
- by (simp only: k_m [symmetric] k_n [symmetric])
- then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
- by (simp add: mult_ac)
- with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
- by simp
- with prod_gcd_lcm [of m n]
- have "lcm m n * gcd q p * gcd m n = k * gcd m n"
- by (simp add: mult_ac)
- with pos_gcd have "lcm m n * gcd q p = k" by simp
- then show ?thesis using dvd_def by auto
-qed
-
-lemma lcm_dvd1 [iff]:
- "m dvd lcm m n"
-proof (cases m)
- case 0 then show ?thesis by simp
-next
- case (Suc _)
- then have mpos: "m > 0" by simp
- show ?thesis
- proof (cases n)
- case 0 then show ?thesis by simp
- next
- case (Suc _)
- then have npos: "n > 0" by simp
- have "gcd m n dvd n" by simp
- then obtain k where "n = gcd m n * k" using dvd_def by auto
- then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
- also have "\<dots> = m * k" using mpos npos gcd_zero by simp
- finally show ?thesis by (simp add: lcm_def)
- qed
-qed
-
-lemma lcm_dvd2 [iff]:
- "n dvd lcm m n"
-proof (cases n)
- case 0 then show ?thesis by simp
-next
- case (Suc _)
- then have npos: "n > 0" by simp
- show ?thesis
- proof (cases m)
- case 0 then show ?thesis by simp
- next
- case (Suc _)
- then have mpos: "m > 0" by simp
- have "gcd m n dvd m" by simp
- then obtain k where "m = gcd m n * k" using dvd_def by auto
- then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
- also have "\<dots> = n * k" using mpos npos gcd_zero by simp
- finally show ?thesis by (simp add: lcm_def)
- qed
-qed
-
-lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
- by (simp add: gcd_commute)
-
-lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
- apply (subgoal_tac "n = m + (n - m)")
- apply (erule ssubst, rule gcd_add1_eq, simp)
- done
-
-
-subsection {* GCD and LCM on integers *}
-
-definition
- zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
- "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
-
-lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i"
-by (simp add: zgcd_def int_dvd_iff)
-
-lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j"
-by (simp add: zgcd_def int_dvd_iff)
-
-lemma zgcd_pos: "zgcd i j \<ge> 0"
-by (simp add: zgcd_def)
-
-lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
-by (simp add: zgcd_def gcd_zero)
-
-lemma zgcd_commute: "zgcd i j = zgcd j i"
-unfolding zgcd_def by (simp add: gcd_commute)
-
-lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
-unfolding zgcd_def by simp
-
-lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
-unfolding zgcd_def by simp
-
- (* should be solved by algebra*)
-lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
- unfolding zgcd_def
-proof -
- assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
- then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
- from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
- have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
- unfolding dvd_def
- by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
- from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
- unfolding dvd_def by blast
- from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
- then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
- then show ?thesis
- apply (subst abs_dvd_iff [symmetric])
- apply (subst dvd_abs_iff [symmetric])
- apply (unfold dvd_def)
- apply (rule_tac x = "int h'" in exI, simp)
- done
-qed
-
-lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
-
-lemma zgcd_greatest:
- assumes "k dvd m" and "k dvd n"
- shows "k dvd zgcd m n"
-proof -
- let ?k' = "nat \<bar>k\<bar>"
- let ?m' = "nat \<bar>m\<bar>"
- let ?n' = "nat \<bar>n\<bar>"
- from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
- unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
- from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
- unfolding zgcd_def by (simp only: zdvd_int)
- then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
- then show "k dvd zgcd m n" by simp
-qed
-
-lemma div_zgcd_relprime:
- assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
- shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
-proof -
- from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith
- let ?g = "zgcd a b"
- let ?a' = "a div ?g"
- let ?b' = "b div ?g"
- let ?g' = "zgcd ?a' ?b'"
- have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
- have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
- from dvdg dvdg' obtain ka kb ka' kb' where
- kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
- unfolding dvd_def by blast
- then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
- then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
- by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
- zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
- have "?g \<noteq> 0" using nz by simp
- then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
- from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
- with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
- with zgcd_pos show "?g' = 1" by simp
-qed
-
-lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
- apply (frule_tac b = n and a = m in pos_mod_sign)
- apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
- apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
- apply (frule_tac a = m in pos_mod_bound)
- apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
- done
-
-lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
- apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
- apply (auto simp add: linorder_neq_iff zgcd_non_0)
- apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
- done
-
-lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
- by (simp add: zgcd_def abs_if)
-
-lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
- by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
-
-lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
- by (simp add: zgcd_def gcd_1_left)
-
-lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
- by (simp add: zgcd_def gcd_assoc)
-
-lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
- apply (rule zgcd_commute [THEN trans])
- apply (rule zgcd_assoc [THEN trans])
- apply (rule zgcd_commute [THEN arg_cong])
- done
-
-lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
- -- {* addition is an AC-operator *}
-
-lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
- by (simp del: minus_mult_right [symmetric]
- add: minus_mult_right nat_mult_distrib zgcd_def abs_if
- mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
-
-lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
- by (simp add: abs_if zgcd_zmult_distrib2)
-
-lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
- by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
-
-lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
- by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
-
-lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
- by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
-
-
-definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
-
-lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
-
-lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
-
-
-lemma dvd_imp_dvd_zlcm1:
- assumes "k dvd i" shows "k dvd (zlcm i j)"
-proof -
- have "nat(abs k) dvd nat(abs i)" using `k dvd i`
- by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
- thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
-qed
-
-lemma dvd_imp_dvd_zlcm2:
- assumes "k dvd j" shows "k dvd (zlcm i j)"
-proof -
- have "nat(abs k) dvd nat(abs j)" using `k dvd j`
- by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
- thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
-qed
-
-
-lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
-by (case_tac "d <0", simp_all)
-
-lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
-by (case_tac "d<0", simp_all)
-
-(* lcm a b is positive for positive a and b *)
-
-lemma lcm_pos:
- assumes mpos: "m > 0"
- and npos: "n>0"
- shows "lcm m n > 0"
-proof(rule ccontr, simp add: lcm_def gcd_zero)
-assume h:"m*n div gcd m n = 0"
-from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
-hence gcdp: "gcd m n > 0" by simp
-with h
-have "m*n < gcd m n"
- by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
-moreover
-have "gcd m n dvd m" by simp
- with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
- with npos have t1:"gcd m n *n \<le> m*n" by simp
- have "gcd m n \<le> gcd m n*n" using npos by simp
- with t1 have "gcd m n \<le> m*n" by arith
-ultimately show "False" by simp
-qed
-
-lemma zlcm_pos:
- assumes anz: "a \<noteq> 0"
- and bnz: "b \<noteq> 0"
- shows "0 < zlcm a b"
-proof-
- let ?na = "nat (abs a)"
- let ?nb = "nat (abs b)"
- have nap: "?na >0" using anz by simp
- have nbp: "?nb >0" using bnz by simp
- have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
- thus ?thesis by (simp add: zlcm_def)
-qed
-
-lemma zgcd_code [code]:
- "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
- by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
-
-end
--- a/src/HOL/Library/Library.thy Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/Library/Library.thy Tue Sep 01 15:39:33 2009 +0200
@@ -43,11 +43,9 @@
OptionalSugar
Option_ord
Permutation
- Pocklington
Poly_Deriv
Polynomial
Preorder
- Primes
Product_Vector
Quicksort
Quotient
--- a/src/HOL/Library/Pocklington.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1263 +0,0 @@
-(* Title: HOL/Library/Pocklington.thy
- Author: Amine Chaieb
-*)
-
-header {* Pocklington's Theorem for Primes *}
-
-theory Pocklington
-imports Main Primes
-begin
-
-definition modeq:: "nat => nat => nat => bool" ("(1[_ = _] '(mod _'))")
- where "[a = b] (mod p) == ((a mod p) = (b mod p))"
-
-definition modneq:: "nat => nat => nat => bool" ("(1[_ \<noteq> _] '(mod _'))")
- where "[a \<noteq> b] (mod p) == ((a mod p) \<noteq> (b mod p))"
-
-lemma modeq_trans:
- "\<lbrakk> [a = b] (mod p); [b = c] (mod p) \<rbrakk> \<Longrightarrow> [a = c] (mod p)"
- by (simp add:modeq_def)
-
-
-lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \<le> x"
- shows "\<exists>q. x = y + n * q"
-using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast
-
-lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"
-unfolding modeq_def nat_mod_eq_iff ..
-
-(* Lemmas about previously defined terms. *)
-
-lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
- (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
- {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
- moreover
- {assume p0: "p\<noteq>0" "p\<noteq>1"
- {assume H: "?lhs"
- {fix m assume m: "m > 0" "m < p"
- {assume "m=1" hence "coprime p m" by simp}
- moreover
- {assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
- have "coprime p m" by simp}
- ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
- hence ?rhs using p0 by auto}
- moreover
- { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
- from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast
- from prime_ge_2[OF q(1)] have q0: "q > 0" by arith
- from dvd_imp_le[OF q(2)] p0 have qp: "q \<le> p" by arith
- {assume "q = p" hence ?lhs using q(1) by blast}
- moreover
- {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
- from H[rule_format, of q] qplt q0 have "coprime p q" by arith
- with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
- ultimately have ?lhs by blast}
- ultimately have ?thesis by blast}
- ultimately show ?thesis by (cases"p=0 \<or> p=1", auto)
-qed
-
-lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
-proof-
- have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
- thus ?thesis by simp
-qed
-
-lemma coprime_mod: assumes n: "n \<noteq> 0" shows "coprime (a mod n) n \<longleftrightarrow> coprime a n"
- using n dvd_mod_iff[of _ n a] by (auto simp add: coprime)
-
-(* Congruences. *)
-
-lemma cong_mod_01[simp,presburger]:
- "[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"
- by (simp_all add: modeq_def, presburger)
-
-lemma cong_sub_cases:
- "[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
-apply (auto simp add: nat_mod)
-apply (rule_tac x="q2" in exI)
-apply (rule_tac x="q1" in exI, simp)
-apply (rule_tac x="q2" in exI)
-apply (rule_tac x="q1" in exI, simp)
-apply (rule_tac x="q1" in exI)
-apply (rule_tac x="q2" in exI, simp)
-apply (rule_tac x="q1" in exI)
-apply (rule_tac x="q2" in exI, simp)
-done
-
-lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)"
- shows "[x = y] (mod n)"
-proof-
- {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) }
- moreover
- {assume az: "a\<noteq>0"
- {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
- with axy cong_sub_cases[of "a*x" "a*y" n] have "[a*(y - x) = 0] (mod n)"
- by (simp only: if_True diff_mult_distrib2)
- hence th: "n dvd a*(y -x)" by simp
- from coprime_divprod[OF th] an have "n dvd y - x"
- by (simp add: coprime_commute)
- hence ?thesis using xy cong_sub_cases[of x y n] by simp}
- moreover
- {assume H: "\<not>x \<le> y" hence xy: "y \<le> x" by arith
- from H az have axy': "\<not> a*x \<le> a*y" by auto
- with axy H cong_sub_cases[of "a*x" "a*y" n] have "[a*(x - y) = 0] (mod n)"
- by (simp only: if_False diff_mult_distrib2)
- hence th: "n dvd a*(x - y)" by simp
- from coprime_divprod[OF th] an have "n dvd x - y"
- by (simp add: coprime_commute)
- hence ?thesis using xy cong_sub_cases[of x y n] by simp}
- ultimately have ?thesis by blast}
- ultimately show ?thesis by blast
-qed
-
-lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)"
- shows "[x = y] (mod n)"
- using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] .
-
-lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def)
-
-lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)
-
-lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"
- by (auto simp add: modeq_def)
-
-lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"
- by (simp add: modeq_def)
-
-lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
- shows "[x + y = x' + y'] (mod n)"
-proof-
- have "(x + y) mod n = (x mod n + y mod n) mod n"
- by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])
- also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp
- also have "\<dots> = (x' + y') mod n"
- by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])
- finally show ?thesis unfolding modeq_def .
-qed
-
-lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
- shows "[x * y = x' * y'] (mod n)"
-proof-
- have "(x * y) mod n = (x mod n) * (y mod n) mod n"
- by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])
- also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp
- also have "\<dots> = (x' * y') mod n"
- by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])
- finally show ?thesis unfolding modeq_def .
-qed
-
-lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
- by (induct k, auto simp add: cong_refl cong_mult)
-lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)"
- and yx: "y <= x" and yx': "y' <= x'"
- shows "[x - y = x' - y'] (mod n)"
-proof-
- { fix x a x' a' y b y' b'
- have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'
- \<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}
- note th = this
- from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"
- and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+
- from th[OF q12 q12' yx yx']
- have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"
- by (simp add: right_distrib)
- thus ?thesis unfolding nat_mod by blast
-qed
-
-lemma cong_mult_lcancel_eq: assumes an: "coprime a n"
- shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .
-next
- assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute)
- from cong_mult_rcancel[OF an H'] show ?rhs .
-qed
-
-lemma cong_mult_rcancel_eq: assumes an: "coprime a n"
- shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)
-
-lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: nat_mod)
-
-lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"
- shows "x = y"
- using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .
-
-lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"
- apply (auto simp add: nat_mod dvd_def)
- apply (rule_tac x="k*q1" in exI)
- apply (rule_tac x="k*q2" in exI)
- by simp
-
-lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp
-
-lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"
- apply (cases "x\<le>1", simp_all)
- using cong_sub_cases[of x 1 n] by auto
-
-lemma cong_divides: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
-apply (auto simp add: nat_mod dvd_def)
-apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
-apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
-done
-
-lemma cong_coprime: assumes xy: "[x = y] (mod n)"
- shows "coprime n x \<longleftrightarrow> coprime n y"
-proof-
- {assume "n=0" hence ?thesis using xy by simp}
- moreover
- {assume nz: "n \<noteq> 0"
- have "coprime n x \<longleftrightarrow> coprime (x mod n) n"
- by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])
- also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp
- also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])
- finally have ?thesis by (simp add: coprime_commute) }
-ultimately show ?thesis by blast
-qed
-
-lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)
-
-lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)
- \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
- by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)
-
-lemma cong_mod_mult: "[x = y] (mod n) \<Longrightarrow> m dvd n \<Longrightarrow> [x = y] (mod m)"
- apply (auto simp add: nat_mod dvd_def)
- apply (rule_tac x="k*q1" in exI)
- apply (rule_tac x="k*q2" in exI, simp)
- done
-
-(* Some things when we know more about the order. *)
-
-lemma cong_le: "y <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
- using nat_mod_lemma[of x y n]
- apply auto
- apply (simp add: nat_mod)
- apply (rule_tac x="q" in exI)
- apply (rule_tac x="q + q" in exI)
- by (auto simp: algebra_simps)
-
-lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
-proof-
- {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
- apply (cases "n=0", simp_all add: cong_commute)
- apply (cases "n=1", simp_all add: cong_commute modeq_def)
- apply arith
- by (cases "a=1", simp_all add: modeq_def cong_commute)}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
- hence ?thesis using cong_le[OF a', of n] by auto }
- ultimately show ?thesis by auto
-qed
-
-(* Some basic theorems about solving congruences. *)
-
-
-lemma cong_solve: assumes an: "coprime a n" shows "\<exists>x. [a * x = b] (mod n)"
-proof-
- {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
- moreover
- {assume az: "a\<noteq>0"
- from bezout_add_strong[OF az, of n]
- obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast
- from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast
- hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp
- hence "a*(x*b) = n*(y*b) + b" by algebra
- hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp
- hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq)
- hence "[a*(x*b) = b] (mod n)" unfolding modeq_def .
- hence ?thesis by blast}
-ultimately show ?thesis by blast
-qed
-
-lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \<noteq> 0"
- shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"
-proof-
- let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"
- from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast
- let ?x = "x mod n"
- from x have th: "[a * ?x = b] (mod n)"
- by (simp add: modeq_def mod_mult_right_eq[of a x n])
- from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp
- {fix y assume Py: "y < n" "[a * y = b] (mod n)"
- from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def)
- hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an])
- with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz
- have "y = ?x" by (simp add: modeq_def)}
- with Px show ?thesis by blast
-qed
-
-lemma cong_solve_unique_nontrivial:
- assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p"
- shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"
-proof-
- from p have p1: "p > 1" using prime_ge_2[OF p] by arith
- hence p01: "p \<noteq> 0" "p \<noteq> 1" by arith+
- from pa have ap: "coprime a p" by (simp add: coprime_commute)
- from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"
- by (auto simp add: coprime_commute)
- from cong_solve_unique[OF px p01(1)]
- obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast
- {assume y0: "y = 0"
- with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])
- with p coprime_prime[OF pa, of p] have False by simp}
- with y show ?thesis unfolding Ex1_def using neq0_conv by blast
-qed
-lemma cong_unique_inverse_prime:
- assumes p: "prime p" and x0: "0 < x" and xp: "x < p"
- shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"
- using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] .
-
-(* Forms of the Chinese remainder theorem. *)
-
-lemma cong_chinese:
- assumes ab: "coprime a b" and xya: "[x = y] (mod a)"
- and xyb: "[x = y] (mod b)"
- shows "[x = y] (mod a*b)"
- using ab xya xyb
- by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]
- cong_sub_cases[of x y "a*b"])
-(cases "x \<le> y", simp_all add: divides_mul[of a _ b])
-
-lemma chinese_remainder_unique:
- assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b\<noteq>0"
- shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
-proof-
- from az bz have abpos: "a*b > 0" by simp
- from chinese_remainder[OF ab az bz] obtain x q1 q2 where
- xq12: "x = m + q1 * a" "x = n + q2 * b" by blast
- let ?w = "x mod (a*b)"
- have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
- from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
- also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
- apply (subst mod_add_left_eq)
- by simp
- finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
- from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
- also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
- also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
- apply (subst mod_add_left_eq)
- by simp
- finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
- {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
- with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
- by (simp_all add: modeq_def)
- from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]
- have "y = ?w" by (simp add: modeq_def)}
- with th1 th2 wab show ?thesis by blast
-qed
-
-lemma chinese_remainder_coprime_unique:
- assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
- and ma: "coprime m a" and nb: "coprime n b"
- shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
-proof-
- let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
- from chinese_remainder_unique[OF ab az bz]
- obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"
- "\<forall>y. ?P y \<longrightarrow> y = x" by blast
- from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]
- have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)
- with coprime_mul[of x a b] have "coprime x (a*b)" by simp
- with x show ?thesis by blast
-qed
-
-(* Euler totient function. *)
-
-definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
-
-lemma phi_0[simp]: "\<phi> 0 = 0"
- unfolding phi_def by auto
-
-lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
-proof-
- have "{ m. 0 < m \<and> m <= n \<and> coprime m n } \<subseteq> {0..n}" by auto
- thus ?thesis by (auto intro: finite_subset)
-qed
-
-declare coprime_1[presburger]
-lemma phi_1[simp]: "\<phi> 1 = 1"
-proof-
- {fix m
- have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }
- thus ?thesis by (simp add: phi_def)
-qed
-
-lemma [simp]: "\<phi> (Suc 0) = Suc 0" using phi_1 by simp
-
-lemma phi_alt: "\<phi>(n) = card { m. coprime m n \<and> m < n}"
-proof-
- {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=0", simp_all)}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- {fix m
- from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
- apply (cases "m = 0", simp_all)
- apply (cases "m = 1", simp_all)
- apply (cases "m = n", auto)
- done }
- hence ?thesis unfolding phi_def by simp}
- ultimately show ?thesis by auto
-qed
-
-lemma phi_finite_lemma[simp]: "finite {m. coprime m n \<and> m < n}" (is "finite ?S")
- by (rule finite_subset[of "?S" "{0..n}"], auto)
-
-lemma phi_another: assumes n: "n\<noteq>1"
- shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
-proof-
- {fix m
- from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
- by (cases "m=0", auto)}
- thus ?thesis unfolding phi_alt by auto
-qed
-
-lemma phi_limit: "\<phi> n \<le> n"
-proof-
- have "{ m. coprime m n \<and> m < n} \<subseteq> {0 ..<n}" by auto
- with card_mono[of "{0 ..<n}" "{ m. coprime m n \<and> m < n}"]
- show ?thesis unfolding phi_alt by auto
-qed
-
-lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
- by auto
-
-lemma phi_limit_strong: assumes n: "n\<noteq>1"
- shows "\<phi>(n) \<le> n - 1"
-proof-
- show ?thesis
- unfolding phi_another[OF n] finite_number_segment[of n, symmetric]
- by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
-qed
-
-lemma phi_lowerbound_1_strong: assumes n: "n \<ge> 1"
- shows "\<phi>(n) \<ge> 1"
-proof-
- let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
- from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt
- apply auto
- apply (cases "n=1", simp_all)
- apply (rule exI[where x=1], simp)
- done
- thus ?thesis by arith
-qed
-
-lemma phi_lowerbound_1: "2 <= n ==> 1 <= \<phi>(n)"
- using phi_lowerbound_1_strong[of n] by auto
-
-lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"
-proof-
- let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
- have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
- by (auto simp add: coprime_commute)
- from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
- from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
- unfolding phi_def by auto
-qed
-
-lemma phi_prime: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1 \<longleftrightarrow> prime n"
-proof-
- {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=1", simp_all)}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- let ?S = "{m. 0 < m \<and> m < n}"
- have fS: "finite ?S" by simp
- let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
- have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto
- {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
- hence ceq: "card ?S' = card ?S"
- using n finite_number_segment[of n] phi_another[OF n(2)] by simp
- {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
- hence mS': "m \<notin> ?S'" by auto
- have "insert m ?S' \<le> ?S" using m by auto
- from m have "card (insert m ?S') \<le> card ?S"
- by - (rule card_mono[of ?S "insert m ?S'"], auto)
- hence False
- unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
- by simp }
- hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast
- hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
- moreover
- {assume H: "prime n"
- hence "?S = ?S'" unfolding prime using n
- by (auto simp add: coprime_commute)
- hence "card ?S = card ?S'" by simp
- hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
- ultimately have ?thesis using n by blast}
- ultimately show ?thesis by (cases "n=0") blast+
-qed
-
-(* Multiplicativity property. *)
-
-lemma phi_multiplicative: assumes ab: "coprime a b"
- shows "\<phi> (a * b) = \<phi> a * \<phi> b"
-proof-
- {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
- hence ?thesis
- by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }
- moreover
- {assume a: "a\<noteq>0" "a\<noteq>1" and b: "b\<noteq>0" "b\<noteq>1"
- hence ab0: "a*b \<noteq> 0" by simp
- let ?S = "\<lambda>k. {m. coprime m k \<and> m < k}"
- let ?f = "\<lambda>x. (x mod a, x mod b)"
- have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"
- proof-
- {fix x assume x:"x \<in> ?S (a*b)"
- hence x': "coprime x (a*b)" "x < a*b" by simp_all
- hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
- from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
- from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
- by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
- moreover
- {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
- hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
- from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
- obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
- "[z = y] (mod b)" by blast
- hence "(x,y) \<in> ?f ` (?S (a*b))"
- using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
- by (auto simp add: image_iff modeq_def)}
- ultimately show ?thesis by auto
- qed
- have finj: "inj_on ?f (?S (a*b))"
- unfolding inj_on_def
- proof(clarify)
- fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
- "y < a * b" "x mod a = y mod a" "x mod b = y mod b"
- hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
- by (simp_all add: coprime_mul_eq)
- from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
- show "x = y" unfolding modeq_def by blast
- qed
- from card_image[OF finj, unfolded eq] have ?thesis
- unfolding phi_alt by simp }
- ultimately show ?thesis by auto
-qed
-
-(* Fermat's Little theorem / Fermat-Euler theorem. *)
-
-
-lemma nproduct_mod:
- assumes fS: "finite S" and n0: "n \<noteq> 0"
- shows "[setprod (\<lambda>m. a(m) mod n) S = setprod a S] (mod n)"
-proof-
- have th1:"[1 = 1] (mod n)" by (simp add: modeq_def)
- from cong_mult
- have th3:"\<forall>x1 y1 x2 y2.
- [x1 = x2] (mod n) \<and> [y1 = y2] (mod n) \<longrightarrow> [x1 * y1 = x2 * y2] (mod n)"
- by blast
- have th4:"\<forall>x\<in>S. [a x mod n = a x] (mod n)" by (simp add: modeq_def)
- from fold_image_related[where h="(\<lambda>m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis unfolding setprod_def by (simp add: fS)
-qed
-
-lemma nproduct_cmul:
- assumes fS:"finite S"
- shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
-unfolding setprod_timesf setprod_constant[OF fS, of c] ..
-
-lemma coprime_nproduct:
- assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
- shows "coprime n (setprod a S)"
- using fS unfolding setprod_def by (rule finite_subset_induct)
- (insert Sn, auto simp add: coprime_mul)
-
-lemma fermat_little: assumes an: "coprime a n"
- shows "[a ^ (\<phi> n) = 1] (mod n)"
-proof-
- {assume "n=0" hence ?thesis by simp}
- moreover
- {assume "n=1" hence ?thesis by (simp add: modeq_def)}
- moreover
- {assume nz: "n \<noteq> 0" and n1: "n \<noteq> 1"
- let ?S = "{m. coprime m n \<and> m < n}"
- let ?P = "\<Prod> ?S"
- have fS: "finite ?S" by simp
- have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..
- {fix m assume m: "m \<in> ?S"
- hence "coprime m n" by simp
- with coprime_mul[of n a m] an have "coprime (a*m) n"
- by (simp add: coprime_commute)}
- hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
- from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
- by (simp add: coprime_commute)
- have Paphi: "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
- proof-
- let ?h = "\<lambda>m. m mod n"
- {fix m assume mS: "m\<in> ?S"
- hence "?h m \<in> ?S" by simp}
- hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
- have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
- hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
-
- have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
- fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
- proof (rule fold_image_eq_general[where h="?h o (op * a)"])
- show "finite ?S" using fS .
- next
- {fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
- from cong_solve_unique[OF an nz, of y]
- obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
- from cong_coprime[OF x(2)] y(1)
- have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
- {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
- hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
- from x(3)[rule_format, of z] z(2,3) have "z=x"
- unfolding modeq_def mod_less[OF y(2)] by simp}
- with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
- unfolding modeq_def mod_less[OF y(2)] by auto }
- thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
- \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
- next
- {fix x assume xS: "x\<in> ?S"
- hence x: "coprime x n" "x < n" by simp_all
- with an have "coprime (a*x) n"
- by (simp add: coprime_mul_eq[of n a x] coprime_commute)
- hence "?h (a*x) \<in> ?S" using nz
- by (simp add: coprime_mod[OF nz] mod_less_divisor)}
- thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
- ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
- ((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp
- qed
- from nproduct_mod[OF fS nz, of "op * a"]
- have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"
- unfolding o_def
- by (simp add: cong_commute)
- also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
- using eq0 fS an by (simp add: setprod_def modeq_def o_def)
- finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
- unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
- nproduct_cmul[OF fS, symmetric] mult_1_right by simp
- qed
- from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
- ultimately show ?thesis by blast
-qed
-
-lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p"
- shows "[a^ (p - 1) = 1] (mod p)"
- using fermat_little[OF ap] p[unfolded phi_prime[symmetric]]
-by simp
-
-
-(* Lucas's theorem. *)
-
-lemma lucas_coprime_lemma:
- assumes m: "m\<noteq>0" and am: "[a^m = 1] (mod n)"
- shows "coprime a n"
-proof-
- {assume "n=1" hence ?thesis by simp}
- moreover
- {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- from m obtain m' where m': "m = Suc m'" by (cases m, blast+)
- {fix d
- assume d: "d dvd a" "d dvd n"
- from n have n1: "1 < n" by arith
- from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp
- from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')
- from dvd_mod_iff[OF d(2), of "a^m"] dam am1
- have "d = 1" by simp }
- hence ?thesis unfolding coprime by auto
- }
- ultimately show ?thesis by blast
-qed
-
-lemma lucas_weak:
- assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"
- and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"
- shows "prime n"
-proof-
- from n have n1: "n \<noteq> 1" "n\<noteq>0" "n - 1 \<noteq> 0" "n - 1 > 0" "n - 1 < n" by arith+
- from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" .
- from fermat_little[OF can] have afn: "[a ^ \<phi> n = 1] (mod n)" .
- {assume "\<phi> n \<noteq> n - 1"
- with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n]
- have c:"\<phi> n > 0 \<and> \<phi> n < n - 1" by arith
- from nm[rule_format, OF c] afn have False ..}
- hence "\<phi> n = n - 1" by blast
- with phi_prime[of n] n1(1,2) show ?thesis by simp
-qed
-
-lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"
- (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume ?rhs thus ?lhs by blast
-next
- assume H: ?lhs then obtain n where n: "P n" by blast
- let ?x = "Least P"
- {fix m assume m: "m < ?x"
- from not_less_Least[OF m] have "\<not> P m" .}
- with LeastI_ex[OF H] show ?rhs by blast
-qed
-
-lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"
- (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
- {assume ?rhs hence ?lhs by blast}
- moreover
- { assume H: ?lhs then obtain n where n: "P n" by blast
- let ?x = "Least P"
- {fix m assume m: "m < ?x"
- from not_less_Least[OF m] have "\<not> P m" .}
- with LeastI_ex[OF H] have ?rhs by blast}
- ultimately show ?thesis by blast
-qed
-
-lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"
-proof(induct n)
- case 0 thus ?case by simp
-next
- case (Suc n)
- have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"
- by (simp add: mod_mult_right_eq[symmetric])
- also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp
- also have "\<dots> = x^(Suc n) mod m"
- by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric])
- finally show ?case .
-qed
-
-lemma lucas:
- assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
- and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"
- shows "prime n"
-proof-
- from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+
- from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp
- from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]
- have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)
- {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
- from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
- m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast
- {assume nm1: "(n - 1) mod m > 0"
- from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
- let ?y = "a^ ((n - 1) div m * m)"
- note mdeq = mod_div_equality[of "(n - 1)" m]
- from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
- of "(n - 1) div m * m"]
- have yn: "coprime ?y n" by (simp add: coprime_commute)
- have "?y mod n = (a^m)^((n - 1) div m) mod n"
- by (simp add: algebra_simps power_mult)
- also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
- using power_mod[of "a^m" n "(n - 1) div m"] by simp
- also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
- by (simp add: power_Suc0)
- finally have th3: "?y mod n = 1" .
- have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
- using an1[unfolded modeq_def onen] onen
- mod_div_equality[of "(n - 1)" m, symmetric]
- by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
- from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
- have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)" .
- from m(4)[rule_format, OF th0] nm1
- less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
- have False by blast }
- hence "(n - 1) mod m = 0" by auto
- then have mn: "m dvd n - 1" by presburger
- then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast
- from n01 r m(2) have r01: "r\<noteq>0" "r\<noteq>1" by - (rule ccontr, simp)+
- from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast
- hence th: "prime p \<and> p dvd n - 1" unfolding r by (auto intro: dvd_mult)
- have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r
- by (simp add: power_mult)
- also have "\<dots> = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp
- also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)
- also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..
- also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)
- finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
- using onen by (simp add: modeq_def)
- with pn[rule_format, OF th] have False by blast}
- hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast
- from lucas_weak[OF n2 an1 th] show ?thesis .
-qed
-
-(* Definition of the order of a number mod n (0 in non-coprime case). *)
-
-definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"
-
-(* This has the expected properties. *)
-
-lemma coprime_ord:
- assumes na: "coprime n a"
- shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
-proof-
- let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
- from euclid[of a] obtain p where p: "prime p" "a < p" by blast
- from na have o: "ord n a = Least ?P" by (simp add: ord_def)
- {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp add: modeq_def)}
- moreover
- {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
- from na have na': "coprime a n" by (simp add: coprime_commute)
- from phi_lowerbound_1[OF n2] fermat_little[OF na']
- have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }
- ultimately have ex: "\<exists>m>0. ?P m" by blast
- from nat_exists_least_iff'[of ?P] ex na show ?thesis
- unfolding o[symmetric] by auto
-qed
-(* With the special value 0 for non-coprime case, it's more convenient. *)
-lemma ord_works:
- "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> ~[a^ m = 1] (mod n))"
-apply (cases "coprime n a")
-using coprime_ord[of n a]
-by (blast, simp add: ord_def modeq_def)
-
-lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast
-lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"
- using ord_works by blast
-lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"
-by (cases "coprime n a", simp add: neq0_conv coprime_ord, simp add: neq0_conv ord_def)
-
-lemma ord_divides:
- "[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume rh: ?rhs
- then obtain k where "d = ord n a * k" unfolding dvd_def by blast
- hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
- by (simp add : modeq_def power_mult power_mod)
- also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
- using ord[of a n, unfolded modeq_def]
- by (simp add: modeq_def power_mod power_Suc0)
- finally show ?lhs .
-next
- assume lh: ?lhs
- { assume H: "\<not> coprime n a"
- hence o: "ord n a = 0" by (simp add: ord_def)
- {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
- moreover
- {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
- from H[unfolded coprime]
- obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto
- from lh[unfolded nat_mod]
- obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
- hence "a ^ d + n * q1 - n * q2 = 1" by simp
- with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
- with p(3) have False by simp
- hence ?rhs ..}
- ultimately have ?rhs by blast}
- moreover
- {assume H: "coprime n a"
- let ?o = "ord n a"
- let ?q = "d div ord n a"
- let ?r = "d mod ord n a"
- from cong_exp[OF ord[of a n], of ?q]
- have eqo: "[(a^?o)^?q = 1] (mod n)" by (simp add: modeq_def power_Suc0)
- from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
- hence op: "?o > 0" by simp
- from mod_div_equality[of d "ord n a"] lh
- have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)
- hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"
- by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])
- hence th: "[a^?r = 1] (mod n)"
- using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
- apply (simp add: modeq_def del: One_nat_def)
- by (simp add: mod_mult_left_eq[symmetric])
- {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
- moreover
- {assume r: "?r \<noteq> 0"
- with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
- from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
- have ?rhs by blast}
- ultimately have ?rhs by blast}
- ultimately show ?rhs by blast
-qed
-
-lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"
-using ord_divides fermat_little coprime_commute by simp
-lemma order_divides_expdiff:
- assumes na: "coprime n a"
- shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
-proof-
- {fix n a d e
- assume na: "coprime n a" and ed: "(e::nat) \<le> d"
- hence "\<exists>c. d = e + c" by arith
- then obtain c where c: "d = e + c" by arith
- from na have an: "coprime a n" by (simp add: coprime_commute)
- from coprime_exp[OF na, of e]
- have aen: "coprime (a^e) n" by (simp add: coprime_commute)
- from coprime_exp[OF na, of c]
- have acn: "coprime (a^c) n" by (simp add: coprime_commute)
- have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
- using c by simp
- also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)
- also have "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"
- using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp
- also have "\<dots> \<longleftrightarrow> ord n a dvd c" by (simp only: ord_divides)
- also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"
- using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides]
- by simp
- finally have "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
- using c by simp }
- note th = this
- have "e \<le> d \<or> d \<le> e" by arith
- moreover
- {assume ed: "e \<le> d" from th[OF na ed] have ?thesis .}
- moreover
- {assume de: "d \<le> e"
- from th[OF na de] have ?thesis by (simp add: cong_commute) }
- ultimately show ?thesis by blast
-qed
-
-(* Another trivial primality characterization. *)
-
-lemma prime_prime_factor:
- "prime n \<longleftrightarrow> n \<noteq> 1\<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"
-proof-
- {assume n: "n=0 \<or> n=1" hence ?thesis using prime_0 two_is_prime by auto}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- {assume pn: "prime n"
-
- from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
- using n
- apply (cases "n = 0 \<or> n=1",simp)
- by (clarsimp, erule_tac x="p" in allE, auto)}
- moreover
- {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
- from n have n1: "n > 1" by arith
- {fix m assume m: "m dvd n" "m\<noteq>1"
- from prime_factor[OF m(2)] obtain p where
- p: "prime p" "p dvd m" by blast
- from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
- with p(2) have "n dvd m" by simp
- hence "m=n" using dvd_anti_sym[OF m(1)] by simp }
- with n1 have "prime n" unfolding prime_def by auto }
- ultimately have ?thesis using n by blast}
- ultimately show ?thesis by auto
-qed
-
-lemma prime_divisor_sqrt:
- "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"
-proof-
- {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
- by (auto simp add: nat_power_eq_0_iff)}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- hence np: "n > 1" by arith
- {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
- from H d have d1n: "d = 1 \<or> d=n" by blast
- {assume dn: "d=n"
- have "n^2 > n*1" using n
- by (simp add: power2_eq_square mult_less_cancel1)
- with dn d(2) have "d=1" by simp}
- with d1n have "d = 1" by blast }
- moreover
- {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
- from d n have "d \<noteq> 0" apply - apply (rule ccontr) by simp
- hence dp: "d > 0" by simp
- from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
- from n dp e have ep:"e > 0" by simp
- have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep
- by (auto simp add: e power2_eq_square mult_le_cancel_left)
- moreover
- {assume h: "d^2 \<le> n"
- from H[rule_format, of d] h d have "d = 1" by blast}
- moreover
- {assume h: "e^2 \<le> n"
- from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
- with H[rule_format, of e] h have "e=1" by simp
- with e have "d = n" by simp}
- ultimately have "d=1 \<or> d=n" by blast}
- ultimately have ?thesis unfolding prime_def using np n(2) by blast}
- ultimately show ?thesis by auto
-qed
-lemma prime_prime_factor_sqrt:
- "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"
- (is "?lhs \<longleftrightarrow>?rhs")
-proof-
- {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
- moreover
- {assume n: "n\<noteq>0" "n\<noteq>1"
- {assume H: ?lhs
- from H[unfolded prime_divisor_sqrt] n
- have ?rhs apply clarsimp by (erule_tac x="p" in allE, simp add: prime_1)
- }
- moreover
- {assume H: ?rhs
- {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
- from prime_factor[OF d(3)]
- obtain p where p: "prime p" "p dvd d" by blast
- from n have np: "n > 0" by arith
- from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
- hence dp: "d > 0" by arith
- from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
- have "p^2 \<le> n" unfolding power2_eq_square by arith
- with H n p(1) dvd_trans[OF p(2) d(1)] have False by blast}
- with n prime_divisor_sqrt have ?lhs by auto}
- ultimately have ?thesis by blast }
- ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)
-qed
-(* Pocklington theorem. *)
-
-lemma pocklington_lemma:
- assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)"
- and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
- and pp: "prime p" and pn: "p dvd n"
- shows "[p = 1] (mod q)"
-proof-
- from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+
- from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]
- obtain k where k: "a ^ (q * r) - 1 = n*k" by blast
- from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast
- {assume a0: "a = 0"
- hence "a^ (n - 1) = 0" using n by (simp add: power_0_left)
- with n an mod_less[of 1 n] have False by (simp add: power_0_left modeq_def)}
- hence a0: "a\<noteq>0" ..
- from n nqr have aqr0: "a ^ (q * r) \<noteq> 0" using a0 by (simp add: neq0_conv)
- hence "(a ^ (q * r) - 1) + 1 = a ^ (q * r)" by simp
- with k l have "a ^ (q * r) = p*l*k + 1" by simp
- hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac)
- hence odq: "ord p (a^r) dvd q"
- unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod by blast
- from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast
- {assume d1: "d \<noteq> 1"
- from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast
- from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
- from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
- from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast
- have P0: "P \<noteq> 0" using P(1) prime_0 by - (rule ccontr, simp)
- from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast
- from d s t P0 have s': "ord p (a^r) * t = s" by algebra
- have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra
- hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
- by (simp only: power_mult)
- have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"
- by (rule cong_exp, rule ord)
- then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
- by (simp add: power_Suc0)
- from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp
- from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp
- with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp
- with p01 pn pd0 have False unfolding coprime by auto}
- hence d1: "d = 1" by blast
- hence o: "ord p (a^r) = q" using d by simp
- from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp
- {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
- from pp[unfolded prime_def] d have dp: "d = p" by blast
- from n have n12:"Suc (n - 2) = n - 1" by arith
- with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]
- have th0: "p dvd a ^ (n - 1)" by simp
- from n have n0: "n \<noteq> 0" by simp
- from d(2) an n12[symmetric] have a0: "a \<noteq> 0"
- by - (rule ccontr, simp add: modeq_def)
- have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by (auto simp add: neq0_conv)
- from coprime_minus1[OF th1, unfolded coprime]
- dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp
- have False by auto}
- hence cpa: "coprime p a" using coprime by auto
- from coprime_exp[OF cpa, of r] coprime_commute
- have arp: "coprime (a^r) p" by blast
- from fermat_little[OF arp, simplified ord_divides] o phip
- have "q dvd (p - 1)" by simp
- then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast
- from prime_0 pp have p0:"p \<noteq> 0" by - (rule ccontr, auto)
- from p0 d have "p + q * 0 = 1 + q * d" by simp
- with nat_mod[of p 1 q, symmetric]
- show ?thesis by blast
-qed
-
-lemma pocklington:
- assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
- and an: "[a^ (n - 1) = 1] (mod n)"
- and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
- shows "prime n"
-unfolding prime_prime_factor_sqrt[of n]
-proof-
- let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"
- from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+
- {fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
- from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)
- hence pq: "p \<le> q" unfolding exp_mono_le .
- from pocklington_lemma[OF n nqr an aq p(1,2)] cong_1_divides
- have th: "q dvd p - 1" by blast
- have "p - 1 \<noteq> 0"using prime_ge_2[OF p(1)] by arith
- with divides_ge[OF th] pq have False by arith }
- with n01 show ?ths by blast
-qed
-
-(* Variant for application, to separate the exponentiation. *)
-lemma pocklington_alt:
- assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
- and an: "[a^ (n - 1) = 1] (mod n)"
- and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)"
- shows "prime n"
-proof-
- {fix p assume p: "prime p" "p dvd q"
- from aq[rule_format] p obtain b where
- b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast
- {assume a0: "a=0"
- from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto
- hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])}
- hence a0: "a\<noteq> 0" ..
- hence a1: "a \<ge> 1" by arith
- from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
- {assume b0: "b = 0"
- from p(2) nqr have "(n - 1) mod p = 0"
- apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
- with mod_div_equality[of "n - 1" p]
- have "(n - 1) div p * p= n - 1" by auto
- hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
- by (simp only: power_mult[symmetric])
- from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
- from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
- from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
- have False by simp}
- then have b0: "b \<noteq> 0" ..
- hence b1: "b \<ge> 1" by arith
- from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr
- have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}
- hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
- by blast
- from pocklington[OF n nqr sqr an th] show ?thesis .
-qed
-
-(* Prime factorizations. *)
-
-definition "primefact ps n = (foldr op * ps 1 = n \<and> (\<forall>p\<in> set ps. prime p))"
-
-lemma primefact: assumes n: "n \<noteq> 0"
- shows "\<exists>ps. primefact ps n"
-using n
-proof(induct n rule: nat_less_induct)
- fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"
- let ?ths = "\<exists>ps. primefact ps n"
- {assume "n = 1"
- hence "primefact [] n" by (simp add: primefact_def)
- hence ?ths by blast }
- moreover
- {assume n1: "n \<noteq> 1"
- with n have n2: "n \<ge> 2" by arith
- from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast
- from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast
- from n m have m0: "m > 0" "m\<noteq>0" by auto
- from prime_ge_2[OF p(1)] have "1 < p" by arith
- with m0 m have mn: "m < n" by auto
- from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" ..
- from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def)
- hence ?ths by blast}
- ultimately show ?ths by blast
-qed
-
-lemma primefact_contains:
- assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"
- shows "p \<in> set ps"
- using pf p pn
-proof(induct ps arbitrary: p n)
- case Nil thus ?case by (auto simp add: primefact_def)
-next
- case (Cons q qs p n)
- from Cons.prems[unfolded primefact_def]
- have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p" and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all
- {assume "p dvd q"
- with p(1) q(1) have "p = q" unfolding prime_def by auto
- hence ?case by simp}
- moreover
- { assume h: "p dvd foldr op * qs 1"
- from q(3) have pqs: "primefact qs (foldr op * qs 1)"
- by (simp add: primefact_def)
- from Cons.hyps[OF pqs p(1) h] have ?case by simp}
- ultimately show ?case using prime_divprod[OF p] by blast
-qed
-
-lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr op * ps 1 = n \<and> list_all prime ps" by (auto simp add: primefact_def list_all_iff)
-
-(* Variant of Lucas theorem. *)
-
-lemma lucas_primefact:
- assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
- and psn: "foldr op * ps 1 = n - 1"
- and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
- shows "prime n"
-proof-
- {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
- from psn psp have psn1: "primefact ps (n - 1)"
- by (auto simp add: list_all_iff primefact_variant)
- from p(3) primefact_contains[OF psn1 p(1,2)] psp
- have False by (induct ps, auto)}
- with lucas[OF n an] show ?thesis by blast
-qed
-
-(* Variant of Pocklington theorem. *)
-
-lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"
-proof-
- from mod_div_equality[of m n]
- have "\<exists>x. x + m mod n = m" by blast
- then show ?thesis by auto
-qed
-
-
-lemma pocklington_primefact:
- assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"
- and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"
- and bqn: "(b^q) mod n = 1"
- and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
- shows "prime n"
-proof-
- from bqn psp qrn
- have bqn: "a ^ (n - 1) mod n = 1"
- and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps" unfolding arnb[symmetric] power_mod
- by (simp_all add: power_mult[symmetric] algebra_simps)
- from n have n0: "n > 0" by arith
- from mod_div_equality[of "a^(n - 1)" n]
- mod_less_divisor[OF n0, of "a^(n - 1)"]
- have an1: "[a ^ (n - 1) = 1] (mod n)"
- unfolding nat_mod bqn
- apply -
- apply (rule exI[where x="0"])
- apply (rule exI[where x="a^(n - 1) div n"])
- by (simp add: algebra_simps)
- {fix p assume p: "prime p" "p dvd q"
- from psp psq have pfpsq: "primefact ps q"
- by (auto simp add: primefact_variant list_all_iff)
- from psp primefact_contains[OF pfpsq p]
- have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
- by (simp add: list_all_iff)
- from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+
- from div_mult1_eq[of r q p] p(2)
- have eq1: "r* (q div p) = (n - 1) div p"
- unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)
- have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith
- from n0 have n00: "n \<noteq> 0" by arith
- from mod_le[OF n00]
- have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .
- {assume "a ^ ((n - 1) div p) mod n = 0"
- then obtain s where s: "a ^ ((n - 1) div p) = n*s"
- unfolding mod_eq_0_iff by blast
- hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
- from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
- from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
- have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
- with eq0 have "a^ (n - 1) = (n*s)^p"
- by (simp add: power_mult[symmetric])
- hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
- also have "\<dots> = 0" by (simp add: mult_assoc)
- finally have False by simp }
- then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
- have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
- unfolding modeq_def by simp
- from cong_sub[OF th1 cong_refl[of 1]] ath[OF th10 th11]
- have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
- by blast
- from cong_coprime[OF th] p'[unfolded eq1]
- have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }
- with pocklington[OF n qrn[symmetric] nq2 an1]
- show ?thesis by blast
-qed
-
-end
--- a/src/HOL/Library/Primes.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,828 +0,0 @@
-(* Title: HOL/Library/Primes.thy
- Author: Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson
- Copyright 1996 University of Cambridge
-*)
-
-header {* Primality on nat *}
-
-theory Primes
-imports Complex_Main Legacy_GCD
-begin
-
-hide (open) const GCD.gcd GCD.coprime GCD.prime
-
-definition
- coprime :: "nat => nat => bool" where
- "coprime m n \<longleftrightarrow> gcd m n = 1"
-
-definition
- prime :: "nat \<Rightarrow> bool" where
- [code del]: "prime p \<longleftrightarrow> (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-
-
-lemma two_is_prime: "prime 2"
- apply (auto simp add: prime_def)
- apply (case_tac m)
- apply (auto dest!: dvd_imp_le)
- done
-
-lemma prime_imp_relprime: "prime p ==> \<not> p dvd n ==> gcd p n = 1"
- apply (auto simp add: prime_def)
- apply (metis One_nat_def gcd_dvd1 gcd_dvd2)
- done
-
-text {*
- This theorem leads immediately to a proof of the uniqueness of
- factorization. If @{term p} divides a product of primes then it is
- one of those primes.
-*}
-
-lemma prime_dvd_mult: "prime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"
- by (blast intro: relprime_dvd_mult prime_imp_relprime)
-
-lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m"
- by (auto dest: prime_dvd_mult)
-
-lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
- by (rule prime_dvd_square) (simp_all add: power2_eq_square)
-
-
-lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0"
-by (induct n, auto)
-
-lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y"
-by(metis linorder_not_less not_less0 power_le_imp_le_base power_less_imp_less_base)
-
-lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y"
-by (simp only: linorder_not_less[symmetric] exp_mono_lt)
-
-lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y"
-using power_inject_base[of x n y] by auto
-
-
-lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n ^ 2 = 4*x"
-proof-
- from e have "2 dvd n" by presburger
- then obtain k where k: "n = 2*k" using dvd_def by auto
- hence "n^2 = 4* (k^2)" by (simp add: power2_eq_square)
- thus ?thesis by blast
-qed
-
-lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n ^ 2 = 4*x + 1"
-proof-
- from e have np: "n > 0" by presburger
- from e have "2 dvd (n - 1)" by presburger
- then obtain k where "n - 1 = 2*k" using dvd_def by auto
- hence k: "n = 2*k + 1" using e by presburger
- hence "n^2 = 4* (k^2 + k) + 1" by algebra
- thus ?thesis by blast
-qed
-
-lemma diff_square: "(x::nat)^2 - y^2 = (x+y)*(x - y)"
-proof-
- have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear)
- moreover
- {assume le: "x \<le> y"
- hence "x ^2 \<le> y^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
- with le have ?thesis by simp }
- moreover
- {assume le: "y \<le> x"
- hence le2: "y ^2 \<le> x^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
- from le have "\<exists>z. y + z = x" by presburger
- then obtain z where z: "x = y + z" by blast
- from le2 have "\<exists>z. x^2 = y^2 + z" by presburger
- then obtain z2 where z2: "x^2 = y^2 + z2" by blast
- from z z2 have ?thesis apply simp by algebra }
- ultimately show ?thesis by blast
-qed
-
-text {* Elementary theory of divisibility *}
-lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto
-lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y"
- using dvd_anti_sym[of x y] by auto
-
-lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)"
- shows "d dvd b"
-proof-
- from da obtain k where k:"a = d*k" by (auto simp add: dvd_def)
- from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def)
- from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2)
- thus ?thesis unfolding dvd_def by blast
-qed
-
-declare nat_mult_dvd_cancel_disj[presburger]
-lemma nat_mult_dvd_cancel_disj'[presburger]:
- "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger
-
-lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"
- by presburger
-
-lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger
-lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m"
- by (auto simp add: dvd_def)
-
-lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)"
-proof(auto simp add: dvd_def)
- fix k assume H: "0 < r" "r < n" "q * n + r = n * k"
- from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute)
- {assume "k - q = 0" with r H(1) have False by simp}
- moreover
- {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
- with H(2) have False by simp}
- ultimately show False by blast
-qed
-lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n"
- by (auto simp add: power_mult_distrib dvd_def)
-
-lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y"
- by (induct n ,auto simp add: dvd_def)
-
-fun fact :: "nat \<Rightarrow> nat" where
- "fact 0 = 1"
-| "fact (Suc n) = Suc n * fact n"
-
-lemma fact_lt: "0 < fact n" by(induct n, simp_all)
-lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp
-lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n"
-proof-
- from le have "\<exists>i. n = m+i" by presburger
- then obtain i where i: "n = m+i" by blast
- have "fact m \<le> fact (m + i)"
- proof(induct m)
- case 0 thus ?case using fact_le[of i] by simp
- next
- case (Suc m)
- have "fact (Suc m) = Suc m * fact m" by simp
- have th1: "Suc m \<le> Suc (m + i)" by simp
- from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps]
- show ?case by simp
- qed
- thus ?thesis using i by simp
-qed
-
-lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n"
-proof(induct n arbitrary: p)
- case 0 thus ?case by simp
-next
- case (Suc n p)
- from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger
- moreover
- {assume "p = Suc n" hence ?case by (simp only: fact.simps dvd_triv_left)}
- moreover
- {assume "p \<le> n"
- with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp
- from dvd_mult[OF th] have ?case by (simp only: fact.simps) }
- ultimately show ?case by blast
-qed
-
-declare dvd_triv_left[presburger]
-declare dvd_triv_right[presburger]
-lemma divides_rexp:
- "x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y])
-
-text {* Coprimality *}
-
-lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
-using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def)
-lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute)
-
-lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)"
-using coprime_def gcd_bezout by auto
-
-lemma coprime_divprod: "d dvd a * b \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
- using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute)
-
-lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def)
-lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def)
-lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def)
-lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def)
-
-lemma gcd_coprime:
- assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b"
- shows "coprime a' b'"
-proof-
- let ?g = "gcd a b"
- {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
- moreover
- {assume az: "a\<noteq> 0"
- from z have z': "?g > 0" by simp
- from bezout_gcd_strong[OF az, of b]
- obtain x y where xy: "a*x = b*y + ?g" by blast
- from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps)
- hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)
- hence "a'*x = (b'*y + 1)"
- by (simp only: nat_mult_eq_cancel1[OF z'])
- hence "a'*x - b'*y = 1" by simp
- with coprime_bezout[of a' b'] have ?thesis by auto}
- ultimately show ?thesis by blast
-qed
-lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def)
-lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b"
- shows "coprime d (a * b)"
-proof-
- from da have th: "gcd a d = 1" by (simp add: coprime_def gcd_commute)
- from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd d (a*b) = 1"
- by (simp add: gcd_commute)
- thus ?thesis unfolding coprime_def .
-qed
-lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b"
-using prems unfolding coprime_bezout
-apply clarsimp
-apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="a*y" in exI)
-apply (simp add: mult_ac)
-apply (rule_tac x="a*x" in exI)
-apply (rule_tac x="y" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a"
-unfolding coprime_bezout
-apply clarsimp
-apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="b*y" in exI)
-apply (simp add: mult_ac)
-apply (rule_tac x="b*x" in exI)
-apply (rule_tac x="y" in exI)
-apply (simp add: mult_ac)
-done
-lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and> coprime d b"
- using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b]
- by blast
-
-lemma gcd_coprime_exists:
- assumes nz: "gcd a b \<noteq> 0"
- shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
-proof-
- let ?g = "gcd a b"
- from gcd_dvd1[of a b] gcd_dvd2[of a b]
- obtain a' b' where "a = ?g*a'" "b = ?g*b'" unfolding dvd_def by blast
- hence ab': "a = a'*?g" "b = b'*?g" by algebra+
- from ab' gcd_coprime[OF nz ab'] show ?thesis by blast
-qed
-
-lemma coprime_exp: "coprime d a ==> coprime d (a^n)"
- by(induct n, simp_all add: coprime_mul)
-
-lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)"
- by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp)
-lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def)
-lemma coprime_plus1[simp]: "coprime (n + 1) n"
- apply (simp add: coprime_bezout)
- apply (rule exI[where x=1])
- apply (rule exI[where x=1])
- apply simp
- done
-lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n"
- using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto
-
-lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd a b ^ n \<or> b ^ n * x - a ^ n * y = gcd a b ^ n"
-proof-
- let ?g = "gcd a b"
- {assume z: "?g = 0" hence ?thesis
- apply (cases n, simp)
- apply arith
- apply (simp only: z power_0_Suc)
- apply (rule exI[where x=0])
- apply (rule exI[where x=0])
- by simp}
- moreover
- {assume z: "?g \<noteq> 0"
- from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where
- ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac)
- hence ab'': "?g*a' = a" "?g * b' = b" by algebra+
- from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n]
- obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1" by blast
- hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n"
- using z by auto
- then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"
- using z ab'' by (simp only: power_mult_distrib[symmetric]
- diff_mult_distrib2 mult_assoc[symmetric])
- hence ?thesis by blast }
- ultimately show ?thesis by blast
-qed
-
-lemma gcd_exp: "gcd (a^n) (b^n) = gcd a b^n"
-proof-
- let ?g = "gcd (a^n) (b^n)"
- let ?gn = "gcd a b^n"
- {fix e assume H: "e dvd a^n" "e dvd b^n"
- from bezout_gcd_pow[of a n b] obtain x y
- where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast
- from dvd_diff_nat [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
- dvd_diff_nat [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
- have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)}
- hence th: "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast
- from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th
- gcd_unique have "?gn = ?g" by blast thus ?thesis by simp
-qed
-
-lemma coprime_exp2: "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b"
-by (simp only: coprime_def gcd_exp exp_eq_1) simp
-
-lemma division_decomp: assumes dc: "(a::nat) dvd b * c"
- shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
-proof-
- let ?g = "gcd a b"
- {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
- apply (rule exI[where x="0"])
- by (rule exI[where x="c"], simp)}
- moreover
- {assume z: "?g \<noteq> 0"
- from gcd_coprime_exists[OF z]
- obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
- from gcd_dvd2[of a b] have thb: "?g dvd b" .
- from ab'(1) have "a' dvd a" unfolding dvd_def by blast
- with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
- from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
- hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
- with z have th_1: "a' dvd b'*c" by simp
- from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" .
- from ab' have "a = ?g*a'" by algebra
- with thb thc have ?thesis by blast }
- ultimately show ?thesis by blast
-qed
-
-lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto)
-
-lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b"
-proof-
- let ?g = "gcd a b"
- from n obtain m where m: "n = Suc m" by (cases n, simp_all)
- {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
- moreover
- {assume z: "?g \<noteq> 0"
- hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
- from gcd_coprime_exists[OF z]
- obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
- from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric])
- hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute)
- with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff)
- have "a' dvd a'^n" by (simp add: m)
- with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
- hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
- from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]]
- have "a' dvd b'" .
- hence "a'*?g dvd b'*?g" by simp
- with ab'(1,2) have ?thesis by simp }
- ultimately show ?thesis by blast
-qed
-
-lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n"
- shows "m * n dvd r"
-proof-
- from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
- unfolding dvd_def by blast
- from mr n' have "m dvd n'*n" by (simp add: mult_commute)
- hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp
- then obtain k where k: "n' = m*k" unfolding dvd_def by blast
- from n' k show ?thesis unfolding dvd_def by auto
-qed
-
-
-text {* A binary form of the Chinese Remainder Theorem. *}
-
-lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0"
- shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
-proof-
- from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a]
- obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"
- and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
- from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified]
- dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto
- let ?x = "v * a * x1 + u * b * x2"
- let ?q1 = "v * x1 + u * y2"
- let ?q2 = "v * y1 + u * x2"
- from dxy2(3)[simplified d12] dxy1(3)[simplified d12]
- have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+
- thus ?thesis by blast
-qed
-
-text {* Primality *}
-
-text {* A few useful theorems about primes *}
-
-lemma prime_0[simp]: "~prime 0" by (simp add: prime_def)
-lemma prime_1[simp]: "~ prime 1" by (simp add: prime_def)
-lemma prime_Suc0[simp]: "~ prime (Suc 0)" by (simp add: prime_def)
-
-lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def)
-lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n"
-using n
-proof(induct n rule: nat_less_induct)
- fix n
- assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1"
- let ?ths = "\<exists>p. prime p \<and> p dvd n"
- {assume "n=0" hence ?ths using two_is_prime by auto}
- moreover
- {assume nz: "n\<noteq>0"
- {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
- moreover
- {assume n: "\<not> prime n"
- with nz H(2)
- obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def)
- from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp
- from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast
- from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast}
- ultimately have ?ths by blast}
- ultimately show ?ths by blast
-qed
-
-lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m"
- shows "m < n"
-proof-
- {assume "m=0" with n have ?thesis by simp}
- moreover
- {assume m: "m \<noteq> 0"
- from npm have mn: "m dvd n" unfolding dvd_def by auto
- from npm m have "n \<noteq> m" using p by auto
- with dvd_imp_le[OF mn] n have ?thesis by simp}
- ultimately show ?thesis by blast
-qed
-
-lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and> p <= Suc (fact n)"
-proof-
- have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith
- from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast
- from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp
- {assume np: "p \<le> n"
- from p(1) have p1: "p \<ge> 1" by (cases p, simp_all)
- from divides_fact[OF p1 np] have pfn': "p dvd fact n" .
- from divides_add_revr[OF pfn' p(2)] p(1) have False by simp}
- hence "n < p" by arith
- with p(1) pfn show ?thesis by auto
-qed
-
-lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto
-
-lemma primes_infinite: "\<not> (finite {p. prime p})"
-apply(simp add: finite_nat_set_iff_bounded_le)
-apply (metis euclid linorder_not_le)
-done
-
-lemma coprime_prime: assumes ab: "coprime a b"
- shows "~(prime p \<and> p dvd a \<and> p dvd b)"
-proof
- assume "prime p \<and> p dvd a \<and> p dvd b"
- thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def)
-qed
-lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))"
- (is "?lhs = ?rhs")
-proof-
- {assume "?lhs" with coprime_prime have ?rhs by blast}
- moreover
- {assume r: "?rhs" and c: "\<not> ?lhs"
- then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast
- from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
- from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)]
- have "p dvd a" "p dvd b" . with p(1) r have False by blast}
- ultimately show ?thesis by blast
-qed
-
-lemma prime_coprime: assumes p: "prime p"
- shows "n = 1 \<or> p dvd n \<or> coprime p n"
-using p prime_imp_relprime[of p n] by (auto simp add: coprime_def)
-
-lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n"
- using prime_coprime[of p n] by auto
-
-declare coprime_0[simp]
-
-lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d])
-lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1"
- shows "\<exists>x y. a * x = b * y + 1"
-proof-
- from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto)
- from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def]
- show ?thesis by auto
-qed
-
-lemma bezout_prime: assumes p: "prime p" and pa: "\<not> p dvd a"
- shows "\<exists>x y. a*x = p*y + 1"
-proof-
- from p have p1: "p \<noteq> 1" using prime_1 by blast
- from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p"
- by (auto simp add: coprime_commute)
- from coprime_bezout_strong[OF ap p1] show ?thesis .
-qed
-lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b"
- shows "p dvd a \<or> p dvd b"
-proof-
- {assume "a=1" hence ?thesis using pab by simp }
- moreover
- {assume "p dvd a" hence ?thesis by blast}
- moreover
- {assume pa: "coprime p a" from coprime_divprod[OF pab pa] have ?thesis .. }
- ultimately show ?thesis using prime_coprime[OF p, of a] by blast
-qed
-
-lemma prime_divprod_eq: assumes p: "prime p"
- shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b"
-using p prime_divprod dvd_mult dvd_mult2 by auto
-
-lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n"
- shows "p dvd x"
-using px
-proof(induct n)
- case 0 thus ?case by simp
-next
- case (Suc n)
- hence th: "p dvd x*x^n" by simp
- {assume H: "p dvd x^n"
- from Suc.hyps[OF H] have ?case .}
- with prime_divprod[OF p th] show ?case by blast
-qed
-
-lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n"
- using prime_divexp[of p x n] divides_exp[of p x n] by blast
-
-lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y"
- shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y"
-proof-
- from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y"
- by blast
- from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
- from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto
-qed
-lemma coprime_sos: assumes xy: "coprime x y"
- shows "coprime (x * y) (x^2 + y^2)"
-proof-
- {assume c: "\<not> coprime (x * y) (x^2 + y^2)"
- from coprime_prime_dvd_ex[OF c] obtain p
- where p: "prime p" "p dvd x*y" "p dvd x^2 + y^2" by blast
- {assume px: "p dvd x"
- from dvd_mult[OF px, of x] p(3)
- obtain r s where "x * x = p * r" and "x^2 + y^2 = p * s"
- by (auto elim!: dvdE)
- then have "y^2 = p * (s - r)"
- by (auto simp add: power2_eq_square diff_mult_distrib2)
- then have "p dvd y^2" ..
- with prime_divexp[OF p(1), of y 2] have py: "p dvd y" .
- from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1
- have False by simp }
- moreover
- {assume py: "p dvd y"
- from dvd_mult[OF py, of y] p(3)
- obtain r s where "y * y = p * r" and "x^2 + y^2 = p * s"
- by (auto elim!: dvdE)
- then have "x^2 = p * (s - r)"
- by (auto simp add: power2_eq_square diff_mult_distrib2)
- then have "p dvd x^2" ..
- with prime_divexp[OF p(1), of x 2] have px: "p dvd x" .
- from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1
- have False by simp }
- ultimately have False using prime_divprod[OF p(1,2)] by blast}
- thus ?thesis by blast
-qed
-
-lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
- unfolding prime_def coprime_prime_eq by blast
-
-lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p"
- shows "coprime x p"
-proof-
- {assume c: "\<not> coprime x p"
- then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast
- from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith
- from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp)
- with g gp p[unfolded prime_def] have False by blast}
-thus ?thesis by blast
-qed
-
-lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger
-lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto
-
-
-text {* One property of coprimality is easier to prove via prime factors. *}
-
-lemma prime_divprod_pow:
- assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b"
- shows "p^n dvd a \<or> p^n dvd b"
-proof-
- {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
- apply (cases "n=0", simp_all)
- apply (cases "a=1", simp_all) done}
- moreover
- {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
- then obtain m where m: "n = Suc m" by (cases n, auto)
- from divides_exp2[OF n pab] have pab': "p dvd a*b" .
- from prime_divprod[OF p pab']
- have "p dvd a \<or> p dvd b" .
- moreover
- {assume pa: "p dvd a"
- have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast
- with prime_coprime[OF p, of b] b
- have cpb: "coprime b p" using coprime_commute by blast
- from coprime_exp[OF cpb] have pnb: "coprime (p^n) b"
- by (simp add: coprime_commute)
- from coprime_divprod[OF pnba pnb] have ?thesis by blast }
- moreover
- {assume pb: "p dvd b"
- have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
- from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast
- with prime_coprime[OF p, of a] a
- have cpb: "coprime a p" using coprime_commute by blast
- from coprime_exp[OF cpb] have pnb: "coprime (p^n) a"
- by (simp add: coprime_commute)
- from coprime_divprod[OF pab pnb] have ?thesis by blast }
- ultimately have ?thesis by blast}
- ultimately show ?thesis by blast
-qed
-
-lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume H: "?lhs"
- hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute)
- thus ?rhs by auto
-next
- assume ?rhs then show ?lhs by auto
-qed
-
-lemma power_Suc0[simp]: "Suc 0 ^ n = Suc 0"
- unfolding One_nat_def[symmetric] power_one ..
-lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n"
- shows "\<exists>r s. a = r^n \<and> b = s ^n"
- using ab abcn
-proof(induct c arbitrary: a b rule: nat_less_induct)
- fix c a b
- assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n"
- let ?ths = "\<exists>r s. a = r^n \<and> b = s ^n"
- {assume n: "n = 0"
- with H(3) power_one have "a*b = 1" by simp
- hence "a = 1 \<and> b = 1" by simp
- hence ?ths
- apply -
- apply (rule exI[where x=1])
- apply (rule exI[where x=1])
- using power_one[of n]
- by simp}
- moreover
- {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
- {assume c: "c = 0"
- with H(3) m H(2) have ?ths apply simp
- apply (cases "a=0", simp_all)
- apply (rule exI[where x="0"], simp)
- apply (rule exI[where x="0"], simp)
- done}
- moreover
- {assume "c=1" with H(3) power_one have "a*b = 1" by simp
- hence "a = 1 \<and> b = 1" by simp
- hence ?ths
- apply -
- apply (rule exI[where x=1])
- apply (rule exI[where x=1])
- using power_one[of n]
- by simp}
- moreover
- {assume c: "c\<noteq>1" "c \<noteq> 0"
- from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast
- from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]]
- have pnab: "p ^ n dvd a \<or> p^n dvd b" .
- from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast
- have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by (simp add: neq0_conv)
- {assume pa: "p^n dvd a"
- then obtain k where k: "a = p^n * k" unfolding dvd_def by blast
- from l have "l dvd c" by auto
- with dvd_imp_le[of l c] c have "l \<le> c" by auto
- moreover {assume "l = c" with l c have "p = 1" by simp with p have False by simp}
- ultimately have lc: "l < c" by arith
- from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]
- have kb: "coprime k b" by (simp add: coprime_commute)
- from H(3) l k pn0 have kbln: "k * b = l ^ n"
- by (auto simp add: power_mult_distrib)
- from H(1)[rule_format, OF lc kb kbln]
- obtain r s where rs: "k = r ^n" "b = s^n" by blast
- from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)
- with rs(2) have ?ths by blast }
- moreover
- {assume pb: "p^n dvd b"
- then obtain k where k: "b = p^n * k" unfolding dvd_def by blast
- from l have "l dvd c" by auto
- with dvd_imp_le[of l c] c have "l \<le> c" by auto
- moreover {assume "l = c" with l c have "p = 1" by simp with p have False by simp}
- ultimately have lc: "l < c" by arith
- from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]
- have kb: "coprime k a" by (simp add: coprime_commute)
- from H(3) l k pn0 n have kbln: "k * a = l ^ n"
- by (simp add: power_mult_distrib mult_commute)
- from H(1)[rule_format, OF lc kb kbln]
- obtain r s where rs: "k = r ^n" "a = s^n" by blast
- from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)
- with rs(2) have ?ths by blast }
- ultimately have ?ths using pnab by blast}
- ultimately have ?ths by blast}
-ultimately show ?ths by blast
-qed
-
-text {* More useful lemmas. *}
-lemma prime_product:
- assumes "prime (p * q)"
- shows "p = 1 \<or> q = 1"
-proof -
- from assms have
- "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
- unfolding prime_def by auto
- from `1 < p * q` have "p \<noteq> 0" by (cases p) auto
- then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
- have "p dvd p * q" by simp
- then have "p = 1 \<or> p = p * q" by (rule P)
- then show ?thesis by (simp add: Q)
-qed
-
-lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1"
-proof(induct n)
- case 0 thus ?case by simp
-next
- case (Suc n)
- {assume "p = 0" hence ?case by simp}
- moreover
- {assume "p=1" hence ?case by simp}
- moreover
- {assume p: "p \<noteq> 0" "p\<noteq>1"
- {assume pp: "prime (p^Suc n)"
- hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp
- with p have n: "n = 0"
- by (simp only: exp_eq_1 ) simp
- with pp have "prime p \<and> Suc n = 1" by simp}
- moreover
- {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
- ultimately have ?case by blast}
- ultimately show ?case by blast
-qed
-
-lemma prime_power_mult:
- assumes p: "prime p" and xy: "x * y = p ^ k"
- shows "\<exists>i j. x = p ^i \<and> y = p^ j"
- using xy
-proof(induct k arbitrary: x y)
- case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
-next
- case (Suc k x y)
- from Suc.prems have pxy: "p dvd x*y" by auto
- from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
- from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)
- {assume px: "p dvd x"
- then obtain d where d: "x = p*d" unfolding dvd_def by blast
- from Suc.prems d have "p*d*y = p^Suc k" by simp
- hence th: "d*y = p^k" using p0 by simp
- from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
- with d have "x = p^Suc i" by simp
- with ij(2) have ?case by blast}
- moreover
- {assume px: "p dvd y"
- then obtain d where d: "y = p*d" unfolding dvd_def by blast
- from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult_commute)
- hence th: "d*x = p^k" using p0 by simp
- from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
- with d have "y = p^Suc i" by simp
- with ij(2) have ?case by blast}
- ultimately show ?case using pxyc by blast
-qed
-
-lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0"
- and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
- using n xn
-proof(induct n arbitrary: k)
- case 0 thus ?case by simp
-next
- case (Suc n k) hence th: "x*x^n = p^k" by simp
- {assume "n = 0" with prems have ?case apply simp
- by (rule exI[where x="k"],simp)}
- moreover
- {assume n: "n \<noteq> 0"
- from prime_power_mult[OF p th]
- obtain i j where ij: "x = p^i" "x^n = p^j"by blast
- from Suc.hyps[OF n ij(2)] have ?case .}
- ultimately show ?case by blast
-qed
-
-lemma divides_primepow: assumes p: "prime p"
- shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
-proof
- assume H: "d dvd p^k" then obtain e where e: "d*e = p^k"
- unfolding dvd_def apply (auto simp add: mult_commute) by blast
- from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
- from prime_ge_2[OF p] have p1: "p > 1" by arith
- from e ij have "p^(i + j) = p^k" by (simp add: power_add)
- hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp
- hence "i \<le> k" by arith
- with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
-next
- {fix i assume H: "i \<le> k" "d = p^i"
- hence "\<exists>j. k = i + j" by arith
- then obtain j where j: "k = i + j" by blast
- hence "p^k = p^j*d" using H(2) by (simp add: power_add)
- hence "d dvd p^k" unfolding dvd_def by auto}
- thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
-qed
-
-lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e"
- by (auto simp add: dvd_def coprime)
-
-declare power_Suc0[simp del]
-declare even_dvd[simp del]
-
-end
--- a/src/HOL/NSA/Examples/NSPrimes.thy Tue Sep 01 14:13:34 2009 +0200
+++ b/src/HOL/NSA/Examples/NSPrimes.thy Tue Sep 01 15:39:33 2009 +0200
@@ -7,7 +7,7 @@
header{*The Nonstandard Primes as an Extension of the Prime Numbers*}
theory NSPrimes
-imports "~~/src/HOL/NumberTheory/Factorization" Hyperreal
+imports "~~/src/HOL/Old_Number_Theory/Factorization" Hyperreal
begin
text{*These can be used to derive an alternative proof of the infinitude of
--- a/src/HOL/NewNumberTheory/Binomial.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,373 +0,0 @@
-(* Title: Binomial.thy
- Authors: Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow
-
-
-Defines the "choose" function, and establishes basic properties.
-
-The original theory "Binomial" was by Lawrence C. Paulson, based on
-the work of Andy Gordon and Florian Kammueller. The approach here,
-which derives the definition of binomial coefficients in terms of the
-factorial function, is due to Jeremy Avigad. The binomial theorem was
-formalized by Tobias Nipkow.
-
-*)
-
-
-header {* Binomial *}
-
-theory Binomial
-imports Cong Fact
-begin
-
-
-subsection {* Main definitions *}
-
-class binomial =
-
-fixes
- binomial :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "choose" 65)
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: binomial
-
-begin
-
-fun
- binomial_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
- "binomial_nat n k =
- (if k = 0 then 1 else
- if n = 0 then 0 else
- (binomial (n - 1) k) + (binomial (n - 1) (k - 1)))"
-
-instance proof qed
-
-end
-
-(* definitions for the integers *)
-
-instantiation int :: binomial
-
-begin
-
-definition
- binomial_int :: "int => int \<Rightarrow> int"
-where
- "binomial_int n k = (if n \<ge> 0 \<and> k \<ge> 0 then int (binomial (nat n) (nat k))
- else 0)"
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-lemma transfer_nat_int_binomial:
- "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow> binomial (nat n) (nat k) =
- nat (binomial n k)"
- unfolding binomial_int_def
- by auto
-
-lemma transfer_nat_int_binomial_closure:
- "n >= (0::int) \<Longrightarrow> k >= 0 \<Longrightarrow> binomial n k >= 0"
- by (auto simp add: binomial_int_def)
-
-declare TransferMorphism_nat_int[transfer add return:
- transfer_nat_int_binomial transfer_nat_int_binomial_closure]
-
-lemma transfer_int_nat_binomial:
- "binomial (int n) (int k) = int (binomial n k)"
- unfolding fact_int_def binomial_int_def by auto
-
-lemma transfer_int_nat_binomial_closure:
- "is_nat n \<Longrightarrow> is_nat k \<Longrightarrow> binomial n k >= 0"
- by (auto simp add: binomial_int_def)
-
-declare TransferMorphism_int_nat[transfer add return:
- transfer_int_nat_binomial transfer_int_nat_binomial_closure]
-
-
-subsection {* Binomial coefficients *}
-
-lemma choose_zero_nat [simp]: "(n::nat) choose 0 = 1"
- by simp
-
-lemma choose_zero_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
- by (simp add: binomial_int_def)
-
-lemma zero_choose_nat [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
- by (induct n rule: induct'_nat, auto)
-
-lemma zero_choose_int [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
- unfolding binomial_int_def apply (case_tac "n < 0")
- apply force
- apply (simp del: binomial_nat.simps)
-done
-
-lemma choose_reduce_nat: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
- (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
- by simp
-
-lemma choose_reduce_int: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
- (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
- unfolding binomial_int_def apply (subst choose_reduce_nat)
- apply (auto simp del: binomial_nat.simps
- simp add: nat_diff_distrib)
-done
-
-lemma choose_plus_one_nat: "((n::nat) + 1) choose (k + 1) =
- (n choose (k + 1)) + (n choose k)"
- by (simp add: choose_reduce_nat)
-
-lemma choose_Suc_nat: "(Suc n) choose (Suc k) =
- (n choose (Suc k)) + (n choose k)"
- by (simp add: choose_reduce_nat One_nat_def)
-
-lemma choose_plus_one_int: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) =
- (n choose (k + 1)) + (n choose k)"
- by (simp add: binomial_int_def choose_plus_one_nat nat_add_distrib del: binomial_nat.simps)
-
-declare binomial_nat.simps [simp del]
-
-lemma choose_self_nat [simp]: "((n::nat) choose n) = 1"
- by (induct n rule: induct'_nat, auto simp add: choose_plus_one_nat)
-
-lemma choose_self_int [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
- by (auto simp add: binomial_int_def)
-
-lemma choose_one_nat [simp]: "(n::nat) choose 1 = n"
- by (induct n rule: induct'_nat, auto simp add: choose_reduce_nat)
-
-lemma choose_one_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
- by (auto simp add: binomial_int_def)
-
-lemma plus_one_choose_self_nat [simp]: "(n::nat) + 1 choose n = n + 1"
- apply (induct n rule: induct'_nat, force)
- apply (case_tac "n = 0")
- apply auto
- apply (subst choose_reduce_nat)
- apply (auto simp add: One_nat_def)
- (* natdiff_cancel_numerals introduces Suc *)
-done
-
-lemma Suc_choose_self_nat [simp]: "(Suc n) choose n = Suc n"
- using plus_one_choose_self_nat by (simp add: One_nat_def)
-
-lemma plus_one_choose_self_int [rule_format, simp]:
- "(n::int) \<ge> 0 \<longrightarrow> n + 1 choose n = n + 1"
- by (auto simp add: binomial_int_def nat_add_distrib)
-
-(* bounded quantification doesn't work with the unicode characters? *)
-lemma choose_pos_nat [rule_format]: "ALL k <= (n::nat).
- ((n::nat) choose k) > 0"
- apply (induct n rule: induct'_nat)
- apply force
- apply clarify
- apply (case_tac "k = 0")
- apply force
- apply (subst choose_reduce_nat)
- apply auto
-done
-
-lemma choose_pos_int: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
- ((n::int) choose k) > 0"
- by (auto simp add: binomial_int_def choose_pos_nat)
-
-lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow>
- (ALL n. P (n + 1) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (k + 1) \<longrightarrow>
- P (n + 1) (k + 1))) \<longrightarrow> (ALL k <= n. P n k)"
- apply (induct n rule: induct'_nat)
- apply auto
- apply (case_tac "k = 0")
- apply auto
- apply (case_tac "k = n + 1")
- apply auto
- apply (drule_tac x = n in spec) back back
- apply (drule_tac x = "k - 1" in spec) back back back
- apply auto
-done
-
-lemma choose_altdef_aux_nat: "(k::nat) \<le> n \<Longrightarrow>
- fact k * fact (n - k) * (n choose k) = fact n"
- apply (rule binomial_induct [of _ k n])
- apply auto
-proof -
- fix k :: nat and n
- assume less: "k < n"
- assume ih1: "fact k * fact (n - k) * (n choose k) = fact n"
- hence one: "fact (k + 1) * fact (n - k) * (n choose k) = (k + 1) * fact n"
- by (subst fact_plus_one_nat, auto)
- assume ih2: "fact (k + 1) * fact (n - (k + 1)) * (n choose (k + 1)) =
- fact n"
- with less have "fact (k + 1) * fact ((n - (k + 1)) + 1) *
- (n choose (k + 1)) = (n - k) * fact n"
- by (subst (2) fact_plus_one_nat, auto)
- with less have two: "fact (k + 1) * fact (n - k) * (n choose (k + 1)) =
- (n - k) * fact n" by simp
- have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
- fact (k + 1) * fact (n - k) * (n choose (k + 1)) +
- fact (k + 1) * fact (n - k) * (n choose k)"
- by (subst choose_reduce_nat, auto simp add: ring_simps)
- also note one
- also note two
- also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)"
- apply (subst fact_plus_one_nat)
- apply (subst left_distrib [symmetric])
- apply simp
- done
- finally show "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
- fact (n + 1)" .
-qed
-
-lemma choose_altdef_nat: "(k::nat) \<le> n \<Longrightarrow>
- n choose k = fact n div (fact k * fact (n - k))"
- apply (frule choose_altdef_aux_nat)
- apply (erule subst)
- apply (simp add: mult_ac)
-done
-
-
-lemma choose_altdef_int:
- assumes "(0::int) <= k" and "k <= n"
- shows "n choose k = fact n div (fact k * fact (n - k))"
-
- apply (subst tsub_eq [symmetric], rule prems)
- apply (rule choose_altdef_nat [transferred])
- using prems apply auto
-done
-
-lemma choose_dvd_nat: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
- unfolding dvd_def apply (frule choose_altdef_aux_nat)
- (* why don't blast and auto get this??? *)
- apply (rule exI)
- apply (erule sym)
-done
-
-lemma choose_dvd_int:
- assumes "(0::int) <= k" and "k <= n"
- shows "fact k * fact (n - k) dvd fact n"
-
- apply (subst tsub_eq [symmetric], rule prems)
- apply (rule choose_dvd_nat [transferred])
- using prems apply auto
-done
-
-(* generalizes Tobias Nipkow's proof to any commutative semiring *)
-theorem binomial: "(a+b::'a::{comm_ring_1,power})^n =
- (SUM k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
-proof (induct n rule: induct'_nat)
- show "?P 0" by simp
-next
- fix n
- assume ih: "?P n"
- have decomp: "{0..n+1} = {0} Un {n+1} Un {1..n}"
- by auto
- have decomp2: "{0..n} = {0} Un {1..n}"
- by auto
- have decomp3: "{1..n+1} = {n+1} Un {1..n}"
- by auto
- have "(a+b)^(n+1) =
- (a+b) * (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
- using ih by (simp add: power_plus_one)
- also have "... = a*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k)) +
- b*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
- by (rule distrib)
- also have "... = (SUM k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
- (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k+1))"
- by (subst (1 2) power_plus_one, simp add: setsum_right_distrib mult_ac)
- also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
- (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
- by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le
- power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
- also have "... = a^(n+1) + b^(n+1) +
- (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
- (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
- by (simp add: decomp2 decomp3)
- also have
- "... = a^(n+1) + b^(n+1) +
- (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
- by (auto simp add: ring_simps setsum_addf [symmetric]
- choose_reduce_nat)
- also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
- using decomp by (simp add: ring_simps)
- finally show "?P (n + 1)" by simp
-qed
-
-lemma set_explicit: "{S. S = T \<and> P S} = (if P T then {T} else {})"
- by auto
-
-lemma card_subsets_nat [rule_format]:
- fixes S :: "'a set"
- assumes "finite S"
- shows "ALL k. card {T. T \<le> S \<and> card T = k} = card S choose k"
- (is "?P S")
-using `finite S`
-proof (induct set: finite)
- show "?P {}" by (auto simp add: set_explicit)
- next fix x :: "'a" and F
- assume iassms: "finite F" "x ~: F"
- assume ih: "?P F"
- show "?P (insert x F)" (is "ALL k. ?Q k")
- proof
- fix k
- show "card {T. T \<subseteq> (insert x F) \<and> card T = k} =
- card (insert x F) choose k" (is "?Q k")
- proof (induct k rule: induct'_nat)
- from iassms have "{T. T \<le> (insert x F) \<and> card T = 0} = {{}}"
- apply auto
- apply (subst (asm) card_0_eq)
- apply (auto elim: finite_subset)
- done
- thus "?Q 0"
- by auto
- next fix k
- show "?Q (k + 1)"
- proof -
- from iassms have fin: "finite (insert x F)" by auto
- hence "{ T. T \<subseteq> insert x F \<and> card T = k + 1} =
- {T. T \<le> F & card T = k + 1} Un
- {T. T \<le> insert x F & x : T & card T = k + 1}"
- by (auto intro!: subsetI)
- with iassms fin have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
- card ({T. T \<subseteq> F \<and> card T = k + 1}) +
- card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1})"
- apply (subst card_Un_disjoint [symmetric])
- apply auto
- (* note: nice! Didn't have to say anything here *)
- done
- also from ih have "card ({T. T \<subseteq> F \<and> card T = k + 1}) =
- card F choose (k+1)" by auto
- also have "card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}) =
- card ({T. T <= F & card T = k})"
- proof -
- let ?f = "%T. T Un {x}"
- from iassms have "inj_on ?f {T. T <= F & card T = k}"
- unfolding inj_on_def by (auto intro!: subsetI)
- hence "card ({T. T <= F & card T = k}) =
- card(?f ` {T. T <= F & card T = k})"
- by (rule card_image [symmetric])
- also from iassms fin have "?f ` {T. T <= F & card T = k} =
- {T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}"
- unfolding image_def
- (* I can't figure out why this next line takes so long *)
- apply auto
- apply (frule (1) finite_subset, force)
- apply (rule_tac x = "xa - {x}" in exI)
- apply (subst card_Diff_singleton)
- apply (auto elim: finite_subset)
- done
- finally show ?thesis by (rule sym)
- qed
- also from ih have "card ({T. T <= F & card T = k}) = card F choose k"
- by auto
- finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
- card F choose (k + 1) + (card F choose k)".
- with iassms choose_plus_one_nat show ?thesis
- by auto
- qed
- qed
- qed
-qed
-
-end
--- a/src/HOL/NewNumberTheory/Cong.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1091 +0,0 @@
-(* Title: HOL/Library/Cong.thy
- ID:
- Authors: Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
- Thomas M. Rasmussen, Jeremy Avigad
-
-
-Defines congruence (notation: [x = y] (mod z)) for natural numbers and
-integers.
-
-This file combines and revises a number of prior developments.
-
-The original theories "GCD" and "Primes" were by Christophe Tabacznyj
-and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
-gcd, lcm, and prime for the natural numbers.
-
-The original theory "IntPrimes" was by Thomas M. Rasmussen, and
-extended gcd, lcm, primes to the integers. Amine Chaieb provided
-another extension of the notions to the integers, and added a number
-of results to "Primes" and "GCD".
-
-The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
-developed the congruence relations on the integers. The notion was
-extended to the natural numbers by Chiaeb. Jeremy Avigad combined
-these, revised and tidied them, made the development uniform for the
-natural numbers and the integers, and added a number of new theorems.
-
-*)
-
-
-header {* Congruence *}
-
-theory Cong
-imports GCD
-begin
-
-subsection {* Turn off One_nat_def *}
-
-lemma induct'_nat [case_names zero plus1, induct type: nat]:
- "\<lbrakk> P (0::nat); !!n. P n \<Longrightarrow> P (n + 1)\<rbrakk> \<Longrightarrow> P n"
-by (erule nat_induct) (simp add:One_nat_def)
-
-lemma cases_nat [case_names zero plus1, cases type: nat]:
- "P (0::nat) \<Longrightarrow> (!!n. P (n + 1)) \<Longrightarrow> P n"
-by(metis induct'_nat)
-
-lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"
-by (simp add: One_nat_def)
-
-lemma power_eq_one_eq_nat [simp]:
- "((x::nat)^m = 1) = (m = 0 | x = 1)"
-by (induct m, auto)
-
-lemma card_insert_if' [simp]: "finite A \<Longrightarrow>
- card (insert x A) = (if x \<in> A then (card A) else (card A) + 1)"
-by (auto simp add: insert_absorb)
-
-(* why wasn't card_insert_if a simp rule? *)
-declare card_insert_disjoint [simp del]
-
-lemma nat_1' [simp]: "nat 1 = 1"
-by simp
-
-(* For those annoying moments where Suc reappears, use Suc_eq_plus1 *)
-
-declare nat_1 [simp del]
-declare add_2_eq_Suc [simp del]
-declare add_2_eq_Suc' [simp del]
-
-
-declare mod_pos_pos_trivial [simp]
-
-
-subsection {* Main definitions *}
-
-class cong =
-
-fixes
- cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(mod _'))")
-
-begin
-
-abbreviation
- notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(mod _'))")
-where
- "notcong x y m == (~cong x y m)"
-
-end
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: cong
-
-begin
-
-definition
- cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
-where
- "cong_nat x y m = ((x mod m) = (y mod m))"
-
-instance proof qed
-
-end
-
-
-(* definitions for the integers *)
-
-instantiation int :: cong
-
-begin
-
-definition
- cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"
-where
- "cong_int x y m = ((x mod m) = (y mod m))"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-
-lemma transfer_nat_int_cong:
- "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> m >= 0 \<Longrightarrow>
- ([(nat x) = (nat y)] (mod (nat m))) = ([x = y] (mod m))"
- unfolding cong_int_def cong_nat_def
- apply (auto simp add: nat_mod_distrib [symmetric])
- apply (subst (asm) eq_nat_nat_iff)
- apply (case_tac "m = 0", force, rule pos_mod_sign, force)+
- apply assumption
-done
-
-declare TransferMorphism_nat_int[transfer add return:
- transfer_nat_int_cong]
-
-lemma transfer_int_nat_cong:
- "[(int x) = (int y)] (mod (int m)) = [x = y] (mod m)"
- apply (auto simp add: cong_int_def cong_nat_def)
- apply (auto simp add: zmod_int [symmetric])
-done
-
-declare TransferMorphism_int_nat[transfer add return:
- transfer_int_nat_cong]
-
-
-subsection {* Congruence *}
-
-(* was zcong_0, etc. *)
-lemma cong_0_nat [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_0_int [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
- by (unfold cong_int_def, auto)
-
-lemma cong_1_nat [simp, presburger]: "[(a::nat) = b] (mod 1)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_Suc_0_nat [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
- by (unfold cong_nat_def, auto simp add: One_nat_def)
-
-lemma cong_1_int [simp, presburger]: "[(a::int) = b] (mod 1)"
- by (unfold cong_int_def, auto)
-
-lemma cong_refl_nat [simp]: "[(k::nat) = k] (mod m)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_refl_int [simp]: "[(k::int) = k] (mod m)"
- by (unfold cong_int_def, auto)
-
-lemma cong_sym_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_sym_int: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
- by (unfold cong_int_def, auto)
-
-lemma cong_sym_eq_nat: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_sym_eq_int: "[(a::int) = b] (mod m) = [b = a] (mod m)"
- by (unfold cong_int_def, auto)
-
-lemma cong_trans_nat [trans]:
- "[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_trans_int [trans]:
- "[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
- by (unfold cong_int_def, auto)
-
-lemma cong_add_nat:
- "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
- apply (unfold cong_nat_def)
- apply (subst (1 2) mod_add_eq)
- apply simp
-done
-
-lemma cong_add_int:
- "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
- apply (unfold cong_int_def)
- apply (subst (1 2) mod_add_left_eq)
- apply (subst (1 2) mod_add_right_eq)
- apply simp
-done
-
-lemma cong_diff_int:
- "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"
- apply (unfold cong_int_def)
- apply (subst (1 2) mod_diff_eq)
- apply simp
-done
-
-lemma cong_diff_aux_int:
- "(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow>
- [c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"
- apply (subst (1 2) tsub_eq)
- apply (auto intro: cong_diff_int)
-done;
-
-lemma cong_diff_nat:
- assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and
- "[c = d] (mod m)"
- shows "[a - c = b - d] (mod m)"
-
- using prems by (rule cong_diff_aux_int [transferred]);
-
-lemma cong_mult_nat:
- "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
- apply (unfold cong_nat_def)
- apply (subst (1 2) mod_mult_eq)
- apply simp
-done
-
-lemma cong_mult_int:
- "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
- apply (unfold cong_int_def)
- apply (subst (1 2) zmod_zmult1_eq)
- apply (subst (1 2) mult_commute)
- apply (subst (1 2) zmod_zmult1_eq)
- apply simp
-done
-
-lemma cong_exp_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
- apply (induct k)
- apply (auto simp add: cong_refl_nat cong_mult_nat)
-done
-
-lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
- apply (induct k)
- apply (auto simp add: cong_refl_int cong_mult_int)
-done
-
-lemma cong_setsum_nat [rule_format]:
- "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
- [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
- apply (case_tac "finite A")
- apply (induct set: finite)
- apply (auto intro: cong_add_nat)
-done
-
-lemma cong_setsum_int [rule_format]:
- "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
- [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
- apply (case_tac "finite A")
- apply (induct set: finite)
- apply (auto intro: cong_add_int)
-done
-
-lemma cong_setprod_nat [rule_format]:
- "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
- [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
- apply (case_tac "finite A")
- apply (induct set: finite)
- apply (auto intro: cong_mult_nat)
-done
-
-lemma cong_setprod_int [rule_format]:
- "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
- [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
- apply (case_tac "finite A")
- apply (induct set: finite)
- apply (auto intro: cong_mult_int)
-done
-
-lemma cong_scalar_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
- by (rule cong_mult_nat, simp_all)
-
-lemma cong_scalar_int: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
- by (rule cong_mult_int, simp_all)
-
-lemma cong_scalar2_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
- by (rule cong_mult_nat, simp_all)
-
-lemma cong_scalar2_int: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
- by (rule cong_mult_int, simp_all)
-
-lemma cong_mult_self_nat: "[(a::nat) * m = 0] (mod m)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_mult_self_int: "[(a::int) * m = 0] (mod m)"
- by (unfold cong_int_def, auto)
-
-lemma cong_eq_diff_cong_0_int: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
- apply (rule iffI)
- apply (erule cong_diff_int [of a b m b b, simplified])
- apply (erule cong_add_int [of "a - b" 0 m b b, simplified])
-done
-
-lemma cong_eq_diff_cong_0_aux_int: "a >= b \<Longrightarrow>
- [(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"
- by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int)
-
-lemma cong_eq_diff_cong_0_nat:
- assumes "(a::nat) >= b"
- shows "[a = b] (mod m) = [a - b = 0] (mod m)"
-
- using prems by (rule cong_eq_diff_cong_0_aux_int [transferred])
-
-lemma cong_diff_cong_0'_nat:
- "[(x::nat) = y] (mod n) \<longleftrightarrow>
- (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
- apply (case_tac "y <= x")
- apply (frule cong_eq_diff_cong_0_nat [where m = n])
- apply auto [1]
- apply (subgoal_tac "x <= y")
- apply (frule cong_eq_diff_cong_0_nat [where m = n])
- apply (subst cong_sym_eq_nat)
- apply auto
-done
-
-lemma cong_altdef_nat: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
- apply (subst cong_eq_diff_cong_0_nat, assumption)
- apply (unfold cong_nat_def)
- apply (simp add: dvd_eq_mod_eq_0 [symmetric])
-done
-
-lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"
- apply (subst cong_eq_diff_cong_0_int)
- apply (unfold cong_int_def)
- apply (simp add: dvd_eq_mod_eq_0 [symmetric])
-done
-
-lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
- by (simp add: cong_altdef_int)
-
-lemma cong_square_int:
- "\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>
- \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
- apply (simp only: cong_altdef_int)
- apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
- (* any way around this? *)
- apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
- apply (auto simp add: ring_simps)
-done
-
-lemma cong_mult_rcancel_int:
- "coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
- apply (subst (1 2) cong_altdef_int)
- apply (subst left_diff_distrib [symmetric])
- apply (rule coprime_dvd_mult_iff_int)
- apply (subst gcd_commute_int, assumption)
-done
-
-lemma cong_mult_rcancel_nat:
- assumes "coprime k (m::nat)"
- shows "[a * k = b * k] (mod m) = [a = b] (mod m)"
-
- apply (rule cong_mult_rcancel_int [transferred])
- using prems apply auto
-done
-
-lemma cong_mult_lcancel_nat:
- "coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"
- by (simp add: mult_commute cong_mult_rcancel_nat)
-
-lemma cong_mult_lcancel_int:
- "coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"
- by (simp add: mult_commute cong_mult_rcancel_int)
-
-(* was zcong_zgcd_zmult_zmod *)
-lemma coprime_cong_mult_int:
- "[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
- \<Longrightarrow> [a = b] (mod m * n)"
- apply (simp only: cong_altdef_int)
- apply (erule (2) divides_mult_int)
-done
-
-lemma coprime_cong_mult_nat:
- assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"
- shows "[a = b] (mod m * n)"
-
- apply (rule coprime_cong_mult_int [transferred])
- using prems apply auto
-done
-
-lemma cong_less_imp_eq_nat: "0 \<le> (a::nat) \<Longrightarrow>
- a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
- by (auto simp add: cong_nat_def mod_pos_pos_trivial)
-
-lemma cong_less_imp_eq_int: "0 \<le> (a::int) \<Longrightarrow>
- a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
- by (auto simp add: cong_int_def mod_pos_pos_trivial)
-
-lemma cong_less_unique_nat:
- "0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
- apply auto
- apply (rule_tac x = "a mod m" in exI)
- apply (unfold cong_nat_def, auto)
-done
-
-lemma cong_less_unique_int:
- "0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
- apply auto
- apply (rule_tac x = "a mod m" in exI)
- apply (unfold cong_int_def, auto simp add: mod_pos_pos_trivial)
-done
-
-lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
- apply (auto simp add: cong_altdef_int dvd_def ring_simps)
- apply (rule_tac [!] x = "-k" in exI, auto)
-done
-
-lemma cong_iff_lin_nat: "([(a::nat) = b] (mod m)) =
- (\<exists>k1 k2. b + k1 * m = a + k2 * m)"
- apply (rule iffI)
- apply (case_tac "b <= a")
- apply (subst (asm) cong_altdef_nat, assumption)
- apply (unfold dvd_def, auto)
- apply (rule_tac x = k in exI)
- apply (rule_tac x = 0 in exI)
- apply (auto simp add: ring_simps)
- apply (subst (asm) cong_sym_eq_nat)
- apply (subst (asm) cong_altdef_nat)
- apply force
- apply (unfold dvd_def, auto)
- apply (rule_tac x = 0 in exI)
- apply (rule_tac x = k in exI)
- apply (auto simp add: ring_simps)
- apply (unfold cong_nat_def)
- apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
- apply (erule ssubst)back
- apply (erule subst)
- apply auto
-done
-
-lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
- apply (subst (asm) cong_iff_lin_int, auto)
- apply (subst add_commute)
- apply (subst (2) gcd_commute_int)
- apply (subst mult_commute)
- apply (subst gcd_add_mult_int)
- apply (rule gcd_commute_int)
-done
-
-lemma cong_gcd_eq_nat:
- assumes "[(a::nat) = b] (mod m)"
- shows "gcd a m = gcd b m"
-
- apply (rule cong_gcd_eq_int [transferred])
- using prems apply auto
-done
-
-lemma cong_imp_coprime_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
- coprime b m"
- by (auto simp add: cong_gcd_eq_nat)
-
-lemma cong_imp_coprime_int: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow>
- coprime b m"
- by (auto simp add: cong_gcd_eq_int)
-
-lemma cong_cong_mod_nat: "[(a::nat) = b] (mod m) =
- [a mod m = b mod m] (mod m)"
- by (auto simp add: cong_nat_def)
-
-lemma cong_cong_mod_int: "[(a::int) = b] (mod m) =
- [a mod m = b mod m] (mod m)"
- by (auto simp add: cong_int_def)
-
-lemma cong_minus_int [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
- by (subst (1 2) cong_altdef_int, auto)
-
-lemma cong_zero_nat [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
- by (auto simp add: cong_nat_def)
-
-lemma cong_zero_int [iff]: "[(a::int) = b] (mod 0) = (a = b)"
- by (auto simp add: cong_int_def)
-
-(*
-lemma mod_dvd_mod_int:
- "0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"
- apply (unfold dvd_def, auto)
- apply (rule mod_mod_cancel)
- apply auto
-done
-
-lemma mod_dvd_mod:
- assumes "0 < (m::nat)" and "m dvd b"
- shows "(a mod b mod m) = (a mod m)"
-
- apply (rule mod_dvd_mod_int [transferred])
- using prems apply auto
-done
-*)
-
-lemma cong_add_lcancel_nat:
- "[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_lcancel_int:
- "[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: cong_iff_lin_int)
-
-lemma cong_add_rcancel_nat: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_rcancel_int: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
- by (simp add: cong_iff_lin_int)
-
-lemma cong_add_lcancel_0_nat: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_lcancel_0_int: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: cong_iff_lin_int)
-
-lemma cong_add_rcancel_0_nat: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_rcancel_0_int: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
- by (simp add: cong_iff_lin_int)
-
-lemma cong_dvd_modulus_nat: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
- [x = y] (mod n)"
- apply (auto simp add: cong_iff_lin_nat dvd_def)
- apply (rule_tac x="k1 * k" in exI)
- apply (rule_tac x="k2 * k" in exI)
- apply (simp add: ring_simps)
-done
-
-lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow>
- [x = y] (mod n)"
- by (auto simp add: cong_altdef_int dvd_def)
-
-lemma cong_dvd_eq_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
- by (unfold cong_nat_def, auto simp add: dvd_eq_mod_eq_0)
-
-lemma cong_dvd_eq_int: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
- by (unfold cong_int_def, auto simp add: dvd_eq_mod_eq_0)
-
-lemma cong_mod_nat: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
- by (simp add: cong_nat_def)
-
-lemma cong_mod_int: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)"
- by (simp add: cong_int_def)
-
-lemma mod_mult_cong_nat: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0
- \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
- by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq)
-
-lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
- apply (simp add: cong_altdef_int)
- apply (subst dvd_minus_iff [symmetric])
- apply (simp add: ring_simps)
-done
-
-lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
- by (auto simp add: cong_altdef_int)
-
-lemma mod_mult_cong_int: "(a::int) ~= 0 \<Longrightarrow> b ~= 0
- \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
- apply (case_tac "b > 0")
- apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)
- apply (subst (1 2) cong_modulus_neg_int)
- apply (unfold cong_int_def)
- apply (subgoal_tac "a * b = (-a * -b)")
- apply (erule ssubst)
- apply (subst zmod_zmult2_eq)
- apply (auto simp add: mod_add_left_eq)
-done
-
-lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
- apply (case_tac "a = 0")
- apply force
- apply (subst (asm) cong_altdef_nat)
- apply auto
-done
-
-lemma cong_0_1_nat: "[(0::nat) = 1] (mod n) = (n = 1)"
- by (unfold cong_nat_def, auto)
-
-lemma cong_0_1_int: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
- by (unfold cong_int_def, auto simp add: zmult_eq_1_iff)
-
-lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \<longleftrightarrow>
- a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
- apply (case_tac "n = 1")
- apply auto [1]
- apply (drule_tac x = "a - 1" in spec)
- apply force
- apply (case_tac "a = 0")
- apply (auto simp add: cong_0_1_nat) [1]
- apply (rule iffI)
- apply (drule cong_to_1_nat)
- apply (unfold dvd_def)
- apply auto [1]
- apply (rule_tac x = k in exI)
- apply (auto simp add: ring_simps) [1]
- apply (subst cong_altdef_nat)
- apply (auto simp add: dvd_def)
-done
-
-lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
- apply (subst cong_altdef_nat)
- apply assumption
- apply (unfold dvd_def, auto simp add: ring_simps)
- apply (rule_tac x = k in exI)
- apply auto
-done
-
-lemma cong_solve_nat: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
- apply (case_tac "n = 0")
- apply force
- apply (frule bezout_nat [of a n], auto)
- apply (rule exI, erule ssubst)
- apply (rule cong_trans_nat)
- apply (rule cong_add_nat)
- apply (subst mult_commute)
- apply (rule cong_mult_self_nat)
- prefer 2
- apply simp
- apply (rule cong_refl_nat)
- apply (rule cong_refl_nat)
-done
-
-lemma cong_solve_int: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
- apply (case_tac "n = 0")
- apply (case_tac "a \<ge> 0")
- apply auto
- apply (rule_tac x = "-1" in exI)
- apply auto
- apply (insert bezout_int [of a n], auto)
- apply (rule exI)
- apply (erule subst)
- apply (rule cong_trans_int)
- prefer 2
- apply (rule cong_add_int)
- apply (rule cong_refl_int)
- apply (rule cong_sym_int)
- apply (rule cong_mult_self_int)
- apply simp
- apply (subst mult_commute)
- apply (rule cong_refl_int)
-done
-
-lemma cong_solve_dvd_nat:
- assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"
- shows "EX x. [a * x = d] (mod n)"
-proof -
- from cong_solve_nat [OF a] obtain x where
- "[a * x = gcd a n](mod n)"
- by auto
- hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
- by (elim cong_scalar2_nat)
- also from b have "(d div gcd a n) * gcd a n = d"
- by (rule dvd_div_mult_self)
- also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
- by auto
- finally show ?thesis
- by auto
-qed
-
-lemma cong_solve_dvd_int:
- assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"
- shows "EX x. [a * x = d] (mod n)"
-proof -
- from cong_solve_int [OF a] obtain x where
- "[a * x = gcd a n](mod n)"
- by auto
- hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)"
- by (elim cong_scalar2_int)
- also from b have "(d div gcd a n) * gcd a n = d"
- by (rule dvd_div_mult_self)
- also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
- by auto
- finally show ?thesis
- by auto
-qed
-
-lemma cong_solve_coprime_nat: "coprime (a::nat) n \<Longrightarrow>
- EX x. [a * x = 1] (mod n)"
- apply (case_tac "a = 0")
- apply force
- apply (frule cong_solve_nat [of a n])
- apply auto
-done
-
-lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow>
- EX x. [a * x = 1] (mod n)"
- apply (case_tac "a = 0")
- apply auto
- apply (case_tac "n \<ge> 0")
- apply auto
- apply (subst cong_int_def, auto)
- apply (frule cong_solve_int [of a n])
- apply auto
-done
-
-lemma coprime_iff_invertible_nat: "m > (1::nat) \<Longrightarrow> coprime a m =
- (EX x. [a * x = 1] (mod m))"
- apply (auto intro: cong_solve_coprime_nat)
- apply (unfold cong_nat_def, auto intro: invertible_coprime_nat)
-done
-
-lemma coprime_iff_invertible_int: "m > (1::int) \<Longrightarrow> coprime a m =
- (EX x. [a * x = 1] (mod m))"
- apply (auto intro: cong_solve_coprime_int)
- apply (unfold cong_int_def)
- apply (auto intro: invertible_coprime_int)
-done
-
-lemma coprime_iff_invertible'_int: "m > (1::int) \<Longrightarrow> coprime a m =
- (EX x. 0 <= x & x < m & [a * x = 1] (mod m))"
- apply (subst coprime_iff_invertible_int)
- apply auto
- apply (auto simp add: cong_int_def)
- apply (rule_tac x = "x mod m" in exI)
- apply (auto simp add: mod_mult_right_eq [symmetric])
-done
-
-
-lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \<Longrightarrow>
- [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
- apply (case_tac "y \<le> x")
- apply (auto simp add: cong_altdef_nat lcm_least_nat) [1]
- apply (rule cong_sym_nat)
- apply (subst (asm) (1 2) cong_sym_eq_nat)
- apply (auto simp add: cong_altdef_nat lcm_least_nat)
-done
-
-lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \<Longrightarrow>
- [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
- by (auto simp add: cong_altdef_int lcm_least_int) [1]
-
-lemma cong_cong_coprime_nat: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
- [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
- apply (frule (1) cong_cong_lcm_nat)back
- apply (simp add: lcm_nat_def)
-done
-
-lemma cong_cong_coprime_int: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
- [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
- apply (frule (1) cong_cong_lcm_int)back
- apply (simp add: lcm_altdef_int cong_abs_int abs_mult [symmetric])
-done
-
-lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \<Longrightarrow>
- (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
- (ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>
- [x = y] (mod (PROD i:A. m i))"
- apply (induct set: finite)
- apply auto
- apply (rule cong_cong_coprime_nat)
- apply (subst gcd_commute_nat)
- apply (rule setprod_coprime_nat)
- apply auto
-done
-
-lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \<Longrightarrow>
- (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
- (ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>
- [x = y] (mod (PROD i:A. m i))"
- apply (induct set: finite)
- apply auto
- apply (rule cong_cong_coprime_int)
- apply (subst gcd_commute_int)
- apply (rule setprod_coprime_int)
- apply auto
-done
-
-lemma binary_chinese_remainder_aux_nat:
- assumes a: "coprime (m1::nat) m2"
- shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
- [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
-proof -
- from cong_solve_coprime_nat [OF a]
- obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
- by auto
- from a have b: "coprime m2 m1"
- by (subst gcd_commute_nat)
- from cong_solve_coprime_nat [OF b]
- obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
- by auto
- have "[m1 * x1 = 0] (mod m1)"
- by (subst mult_commute, rule cong_mult_self_nat)
- moreover have "[m2 * x2 = 0] (mod m2)"
- by (subst mult_commute, rule cong_mult_self_nat)
- moreover note one two
- ultimately show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_aux_int:
- assumes a: "coprime (m1::int) m2"
- shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
- [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
-proof -
- from cong_solve_coprime_int [OF a]
- obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
- by auto
- from a have b: "coprime m2 m1"
- by (subst gcd_commute_int)
- from cong_solve_coprime_int [OF b]
- obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
- by auto
- have "[m1 * x1 = 0] (mod m1)"
- by (subst mult_commute, rule cong_mult_self_int)
- moreover have "[m2 * x2 = 0] (mod m2)"
- by (subst mult_commute, rule cong_mult_self_int)
- moreover note one two
- ultimately show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_nat:
- assumes a: "coprime (m1::nat) m2"
- shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
- from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2
- where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
- "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
- by blast
- let ?x = "u1 * b1 + u2 * b2"
- have "[?x = u1 * 1 + u2 * 0] (mod m1)"
- apply (rule cong_add_nat)
- apply (rule cong_scalar2_nat)
- apply (rule `[b1 = 1] (mod m1)`)
- apply (rule cong_scalar2_nat)
- apply (rule `[b2 = 0] (mod m1)`)
- done
- hence "[?x = u1] (mod m1)" by simp
- have "[?x = u1 * 0 + u2 * 1] (mod m2)"
- apply (rule cong_add_nat)
- apply (rule cong_scalar2_nat)
- apply (rule `[b1 = 0] (mod m2)`)
- apply (rule cong_scalar2_nat)
- apply (rule `[b2 = 1] (mod m2)`)
- done
- hence "[?x = u2] (mod m2)" by simp
- with `[?x = u1] (mod m1)` show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_int:
- assumes a: "coprime (m1::int) m2"
- shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
- from binary_chinese_remainder_aux_int [OF a] obtain b1 b2
- where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
- "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
- by blast
- let ?x = "u1 * b1 + u2 * b2"
- have "[?x = u1 * 1 + u2 * 0] (mod m1)"
- apply (rule cong_add_int)
- apply (rule cong_scalar2_int)
- apply (rule `[b1 = 1] (mod m1)`)
- apply (rule cong_scalar2_int)
- apply (rule `[b2 = 0] (mod m1)`)
- done
- hence "[?x = u1] (mod m1)" by simp
- have "[?x = u1 * 0 + u2 * 1] (mod m2)"
- apply (rule cong_add_int)
- apply (rule cong_scalar2_int)
- apply (rule `[b1 = 0] (mod m2)`)
- apply (rule cong_scalar2_int)
- apply (rule `[b2 = 1] (mod m2)`)
- done
- hence "[?x = u2] (mod m2)" by simp
- with `[?x = u1] (mod m1)` show ?thesis by blast
-qed
-
-lemma cong_modulus_mult_nat: "[(x::nat) = y] (mod m * n) \<Longrightarrow>
- [x = y] (mod m)"
- apply (case_tac "y \<le> x")
- apply (simp add: cong_altdef_nat)
- apply (erule dvd_mult_left)
- apply (rule cong_sym_nat)
- apply (subst (asm) cong_sym_eq_nat)
- apply (simp add: cong_altdef_nat)
- apply (erule dvd_mult_left)
-done
-
-lemma cong_modulus_mult_int: "[(x::int) = y] (mod m * n) \<Longrightarrow>
- [x = y] (mod m)"
- apply (simp add: cong_altdef_int)
- apply (erule dvd_mult_left)
-done
-
-lemma cong_less_modulus_unique_nat:
- "[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
- by (simp add: cong_nat_def)
-
-lemma binary_chinese_remainder_unique_nat:
- assumes a: "coprime (m1::nat) m2" and
- nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
- shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
- from binary_chinese_remainder_nat [OF a] obtain y where
- "[y = u1] (mod m1)" and "[y = u2] (mod m2)"
- by blast
- let ?x = "y mod (m1 * m2)"
- from nz have less: "?x < m1 * m2"
- by auto
- have one: "[?x = u1] (mod m1)"
- apply (rule cong_trans_nat)
- prefer 2
- apply (rule `[y = u1] (mod m1)`)
- apply (rule cong_modulus_mult_nat)
- apply (rule cong_mod_nat)
- using nz apply auto
- done
- have two: "[?x = u2] (mod m2)"
- apply (rule cong_trans_nat)
- prefer 2
- apply (rule `[y = u2] (mod m2)`)
- apply (subst mult_commute)
- apply (rule cong_modulus_mult_nat)
- apply (rule cong_mod_nat)
- using nz apply auto
- done
- have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow>
- z = ?x"
- proof (clarify)
- fix z
- assume "z < m1 * m2"
- assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)"
- have "[?x = z] (mod m1)"
- apply (rule cong_trans_nat)
- apply (rule `[?x = u1] (mod m1)`)
- apply (rule cong_sym_nat)
- apply (rule `[z = u1] (mod m1)`)
- done
- moreover have "[?x = z] (mod m2)"
- apply (rule cong_trans_nat)
- apply (rule `[?x = u2] (mod m2)`)
- apply (rule cong_sym_nat)
- apply (rule `[z = u2] (mod m2)`)
- done
- ultimately have "[?x = z] (mod m1 * m2)"
- by (auto intro: coprime_cong_mult_nat a)
- with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"
- apply (intro cong_less_modulus_unique_nat)
- apply (auto, erule cong_sym_nat)
- done
- qed
- with less one two show ?thesis
- by auto
- qed
-
-lemma chinese_remainder_aux_nat:
- fixes A :: "'a set" and
- m :: "'a \<Rightarrow> nat"
- assumes fin: "finite A" and
- cop: "ALL i : A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
- shows "EX b. (ALL i : A.
- [b i = 1] (mod m i) \<and> [b i = 0] (mod (PROD j : A - {i}. m j)))"
-proof (rule finite_set_choice, rule fin, rule ballI)
- fix i
- assume "i : A"
- with cop have "coprime (PROD j : A - {i}. m j) (m i)"
- by (intro setprod_coprime_nat, auto)
- hence "EX x. [(PROD j : A - {i}. m j) * x = 1] (mod m i)"
- by (elim cong_solve_coprime_nat)
- then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
- by auto
- moreover have "[(PROD j : A - {i}. m j) * x = 0]
- (mod (PROD j : A - {i}. m j))"
- by (subst mult_commute, rule cong_mult_self_nat)
- ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0]
- (mod setprod m (A - {i}))"
- by blast
-qed
-
-lemma chinese_remainder_nat:
- fixes A :: "'a set" and
- m :: "'a \<Rightarrow> nat" and
- u :: "'a \<Rightarrow> nat"
- assumes
- fin: "finite A" and
- cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
- shows "EX x. (ALL i:A. [x = u i] (mod m i))"
-proof -
- from chinese_remainder_aux_nat [OF fin cop] obtain b where
- bprop: "ALL i:A. [b i = 1] (mod m i) \<and>
- [b i = 0] (mod (PROD j : A - {i}. m j))"
- by blast
- let ?x = "SUM i:A. (u i) * (b i)"
- show "?thesis"
- proof (rule exI, clarify)
- fix i
- assume a: "i : A"
- show "[?x = u i] (mod m i)"
- proof -
- from fin a have "?x = (SUM j:{i}. u j * b j) +
- (SUM j:A-{i}. u j * b j)"
- by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)
- hence "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
- by auto
- also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
- u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
- apply (rule cong_add_nat)
- apply (rule cong_scalar2_nat)
- using bprop a apply blast
- apply (rule cong_setsum_nat)
- apply (rule cong_scalar2_nat)
- using bprop apply auto
- apply (rule cong_dvd_modulus_nat)
- apply (drule (1) bspec)
- apply (erule conjE)
- apply assumption
- apply (rule dvd_setprod)
- using fin a apply auto
- done
- finally show ?thesis
- by simp
- qed
- qed
-qed
-
-lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow>
- (ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
- (ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>
- [x = y] (mod (PROD i:A. m i))"
- apply (induct set: finite)
- apply auto
- apply (erule (1) coprime_cong_mult_nat)
- apply (subst gcd_commute_nat)
- apply (rule setprod_coprime_nat)
- apply auto
-done
-
-lemma chinese_remainder_unique_nat:
- fixes A :: "'a set" and
- m :: "'a \<Rightarrow> nat" and
- u :: "'a \<Rightarrow> nat"
- assumes
- fin: "finite A" and
- nz: "ALL i:A. m i \<noteq> 0" and
- cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
- shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"
-proof -
- from chinese_remainder_nat [OF fin cop] obtain y where
- one: "(ALL i:A. [y = u i] (mod m i))"
- by blast
- let ?x = "y mod (PROD i:A. m i)"
- from fin nz have prodnz: "(PROD i:A. m i) \<noteq> 0"
- by auto
- hence less: "?x < (PROD i:A. m i)"
- by auto
- have cong: "ALL i:A. [?x = u i] (mod m i)"
- apply auto
- apply (rule cong_trans_nat)
- prefer 2
- using one apply auto
- apply (rule cong_dvd_modulus_nat)
- apply (rule cong_mod_nat)
- using prodnz apply auto
- apply (rule dvd_setprod)
- apply (rule fin)
- apply assumption
- done
- have unique: "ALL z. z < (PROD i:A. m i) \<and>
- (ALL i:A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"
- proof (clarify)
- fix z
- assume zless: "z < (PROD i:A. m i)"
- assume zcong: "(ALL i:A. [z = u i] (mod m i))"
- have "ALL i:A. [?x = z] (mod m i)"
- apply clarify
- apply (rule cong_trans_nat)
- using cong apply (erule bspec)
- apply (rule cong_sym_nat)
- using zcong apply auto
- done
- with fin cop have "[?x = z] (mod (PROD i:A. m i))"
- by (intro coprime_cong_prod_nat, auto)
- with zless less show "z = ?x"
- apply (intro cong_less_modulus_unique_nat)
- apply (auto, erule cong_sym_nat)
- done
- qed
- from less cong unique show ?thesis
- by blast
-qed
-
-end
--- a/src/HOL/NewNumberTheory/Fib.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,319 +0,0 @@
-(* Title: Fib.thy
- Authors: Lawrence C. Paulson, Jeremy Avigad
-
-
-Defines the fibonacci function.
-
-The original "Fib" is due to Lawrence C. Paulson, and was adapted by
-Jeremy Avigad.
-*)
-
-
-header {* Fib *}
-
-theory Fib
-imports Binomial
-begin
-
-
-subsection {* Main definitions *}
-
-class fib =
-
-fixes
- fib :: "'a \<Rightarrow> 'a"
-
-
-(* definition for the natural numbers *)
-
-instantiation nat :: fib
-
-begin
-
-fun
- fib_nat :: "nat \<Rightarrow> nat"
-where
- "fib_nat n =
- (if n = 0 then 0 else
- (if n = 1 then 1 else
- fib (n - 1) + fib (n - 2)))"
-
-instance proof qed
-
-end
-
-(* definition for the integers *)
-
-instantiation int :: fib
-
-begin
-
-definition
- fib_int :: "int \<Rightarrow> int"
-where
- "fib_int n = (if n >= 0 then int (fib (nat n)) else 0)"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-
-lemma transfer_nat_int_fib:
- "(x::int) >= 0 \<Longrightarrow> fib (nat x) = nat (fib x)"
- unfolding fib_int_def by auto
-
-lemma transfer_nat_int_fib_closure:
- "n >= (0::int) \<Longrightarrow> fib n >= 0"
- by (auto simp add: fib_int_def)
-
-declare TransferMorphism_nat_int[transfer add return:
- transfer_nat_int_fib transfer_nat_int_fib_closure]
-
-lemma transfer_int_nat_fib:
- "fib (int n) = int (fib n)"
- unfolding fib_int_def by auto
-
-lemma transfer_int_nat_fib_closure:
- "is_nat n \<Longrightarrow> fib n >= 0"
- unfolding fib_int_def by auto
-
-declare TransferMorphism_int_nat[transfer add return:
- transfer_int_nat_fib transfer_int_nat_fib_closure]
-
-
-subsection {* Fibonacci numbers *}
-
-lemma fib_0_nat [simp]: "fib (0::nat) = 0"
- by simp
-
-lemma fib_0_int [simp]: "fib (0::int) = 0"
- unfolding fib_int_def by simp
-
-lemma fib_1_nat [simp]: "fib (1::nat) = 1"
- by simp
-
-lemma fib_Suc_0_nat [simp]: "fib (Suc 0) = Suc 0"
- by simp
-
-lemma fib_1_int [simp]: "fib (1::int) = 1"
- unfolding fib_int_def by simp
-
-lemma fib_reduce_nat: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
- by simp
-
-declare fib_nat.simps [simp del]
-
-lemma fib_reduce_int: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
- unfolding fib_int_def
- by (auto simp add: fib_reduce_nat nat_diff_distrib)
-
-lemma fib_neg_int [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
- unfolding fib_int_def by auto
-
-lemma fib_2_nat [simp]: "fib (2::nat) = 1"
- by (subst fib_reduce_nat, auto)
-
-lemma fib_2_int [simp]: "fib (2::int) = 1"
- by (subst fib_reduce_int, auto)
-
-lemma fib_plus_2_nat: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
- by (subst fib_reduce_nat, auto simp add: One_nat_def)
-(* the need for One_nat_def is due to the natdiff_cancel_numerals
- procedure *)
-
-lemma fib_induct_nat: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow>
- (!!n. P n \<Longrightarrow> P (n + 1) \<Longrightarrow> P (n + 2)) \<Longrightarrow> P n"
- apply (atomize, induct n rule: nat_less_induct)
- apply auto
- apply (case_tac "n = 0", force)
- apply (case_tac "n = 1", force)
- apply (subgoal_tac "n >= 2")
- apply (frule_tac x = "n - 1" in spec)
- apply (drule_tac x = "n - 2" in spec)
- apply (drule_tac x = "n - 2" in spec)
- apply auto
- apply (auto simp add: One_nat_def) (* again, natdiff_cancel *)
-done
-
-lemma fib_add_nat: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) +
- fib k * fib n"
- apply (induct n rule: fib_induct_nat)
- apply auto
- apply (subst fib_reduce_nat)
- apply (auto simp add: ring_simps)
- apply (subst (1 3 5) fib_reduce_nat)
- apply (auto simp add: ring_simps Suc_eq_plus1)
-(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
- apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
- apply (erule ssubst) back back
- apply (erule ssubst) back
- apply auto
-done
-
-lemma fib_add'_nat: "fib (n + Suc k) = fib (Suc k) * fib (Suc n) +
- fib k * fib n"
- using fib_add_nat by (auto simp add: One_nat_def)
-
-
-(* transfer from nats to ints *)
-lemma fib_add_int [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
- fib (n + k + 1) = fib (k + 1) * fib (n + 1) +
- fib k * fib n "
-
- by (rule fib_add_nat [transferred])
-
-lemma fib_neq_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
- apply (induct n rule: fib_induct_nat)
- apply (auto simp add: fib_plus_2_nat)
-done
-
-lemma fib_gr_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
- by (frule fib_neq_0_nat, simp)
-
-lemma fib_gr_0_int: "(n::int) > 0 \<Longrightarrow> fib n > 0"
- unfolding fib_int_def by (simp add: fib_gr_0_nat)
-
-text {*
- \medskip Concrete Mathematics, page 278: Cassini's identity. The proof is
- much easier using integers, not natural numbers!
-*}
-
-lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) -
- (fib (int n + 1))^2 = (-1)^(n + 1)"
- apply (induct n)
- apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
- power_add)
-done
-
-lemma fib_Cassini_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n -
- (fib (n + 1))^2 = (-1)^(nat n + 1)"
- by (insert fib_Cassini_aux_int [of "nat n"], auto)
-
-(*
-lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n =
- (fib (n + 1))^2 + (-1)^(nat n + 1)"
- by (frule fib_Cassini_int, simp)
-*)
-
-lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
- (if even n then tsub ((fib (n + 1))^2) 1
- else (fib (n + 1))^2 + 1)"
- apply (frule fib_Cassini_int, auto simp add: pos_int_even_equiv_nat_even)
- apply (subst tsub_eq)
- apply (insert fib_gr_0_int [of "n + 1"], force)
- apply auto
-done
-
-lemma fib_Cassini_nat: "fib ((n::nat) + 2) * fib n =
- (if even n then (fib (n + 1))^2 - 1
- else (fib (n + 1))^2 + 1)"
-
- by (rule fib_Cassini'_int [transferred, of n], auto)
-
-
-text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
-
-lemma coprime_fib_plus_1_nat: "coprime (fib (n::nat)) (fib (n + 1))"
- apply (induct n rule: fib_induct_nat)
- apply auto
- apply (subst (2) fib_reduce_nat)
- apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
- apply (subst add_commute, auto)
- apply (subst gcd_commute_nat, auto simp add: ring_simps)
-done
-
-lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
- using coprime_fib_plus_1_nat by (simp add: One_nat_def)
-
-lemma coprime_fib_plus_1_int:
- "n >= 0 \<Longrightarrow> coprime (fib (n::int)) (fib (n + 1))"
- by (erule coprime_fib_plus_1_nat [transferred])
-
-lemma gcd_fib_add_nat: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
- apply (simp add: gcd_commute_nat [of "fib m"])
- apply (rule cases_nat [of _ m])
- apply simp
- apply (subst add_assoc [symmetric])
- apply (simp add: fib_add_nat)
- apply (subst gcd_commute_nat)
- apply (subst mult_commute)
- apply (subst gcd_add_mult_nat)
- apply (subst gcd_commute_nat)
- apply (rule gcd_mult_cancel_nat)
- apply (rule coprime_fib_plus_1_nat)
-done
-
-lemma gcd_fib_add_int [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
- gcd (fib (m::int)) (fib (n + m)) = gcd (fib m) (fib n)"
- by (erule gcd_fib_add_nat [transferred])
-
-lemma gcd_fib_diff_nat: "(m::nat) \<le> n \<Longrightarrow>
- gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
- by (simp add: gcd_fib_add_nat [symmetric, of _ "n-m"])
-
-lemma gcd_fib_diff_int: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow>
- gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
- by (simp add: gcd_fib_add_int [symmetric, of _ "n-m"])
-
-lemma gcd_fib_mod_nat: "0 < (m::nat) \<Longrightarrow>
- gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-proof (induct n rule: less_induct)
- case (less n)
- from less.prems have pos_m: "0 < m" .
- show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
- proof (cases "m < n")
- case True note m_n = True
- then have m_n': "m \<le> n" by auto
- with pos_m have pos_n: "0 < n" by auto
- with pos_m m_n have diff: "n - m < n" by auto
- have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
- by (simp add: mod_if [of n]) (insert m_n, auto)
- also have "\<dots> = gcd (fib m) (fib (n - m))"
- by (simp add: less.hyps diff pos_m)
- also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff_nat m_n')
- finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
- next
- case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
- by (cases "m = n") auto
- qed
-qed
-
-lemma gcd_fib_mod_int:
- assumes "0 < (m::int)" and "0 <= n"
- shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-
- apply (rule gcd_fib_mod_nat [transferred])
- using prems apply auto
-done
-
-lemma fib_gcd_nat: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"
- -- {* Law 6.111 *}
- apply (induct m n rule: gcd_nat_induct)
- apply (simp_all add: gcd_non_0_nat gcd_commute_nat gcd_fib_mod_nat)
-done
-
-lemma fib_gcd_int: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
- fib (gcd (m::int) n) = gcd (fib m) (fib n)"
- by (erule fib_gcd_nat [transferred])
-
-lemma atMost_plus_one_nat: "{..(k::nat) + 1} = insert (k + 1) {..k}"
- by auto
-
-theorem fib_mult_eq_setsum_nat:
- "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
- apply (induct n)
- apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
-done
-
-theorem fib_mult_eq_setsum'_nat:
- "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
- using fib_mult_eq_setsum_nat by (simp add: One_nat_def)
-
-theorem fib_mult_eq_setsum_int [rule_format]:
- "n >= 0 \<Longrightarrow> fib ((n::int) + 1) * fib n = (\<Sum>k \<in> {0..n}. fib k * fib k)"
- by (erule fib_mult_eq_setsum_nat [transferred])
-
-end
--- a/src/HOL/NewNumberTheory/MiscAlgebra.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,355 +0,0 @@
-(* Title: MiscAlgebra.thy
- Author: Jeremy Avigad
-
-These are things that can be added to the Algebra library.
-*)
-
-theory MiscAlgebra
-imports
- "~~/src/HOL/Algebra/Ring"
- "~~/src/HOL/Algebra/FiniteProduct"
-begin;
-
-(* finiteness stuff *)
-
-lemma bounded_set1_int [intro]: "finite {(x::int). a < x & x < b & P x}"
- apply (subgoal_tac "{x. a < x & x < b & P x} <= {a<..<b}")
- apply (erule finite_subset)
- apply auto
-done
-
-
-(* The rest is for the algebra libraries *)
-
-(* These go in Group.thy. *)
-
-(*
- Show that the units in any monoid give rise to a group.
-
- The file Residues.thy provides some infrastructure to use
- facts about the unit group within the ring locale.
-*)
-
-
-constdefs
- units_of :: "('a, 'b) monoid_scheme => 'a monoid"
- "units_of G == (| carrier = Units G,
- Group.monoid.mult = Group.monoid.mult G,
- one = one G |)";
-
-(*
-
-lemma (in monoid) Units_mult_closed [intro]:
- "x : Units G ==> y : Units G ==> x \<otimes> y : Units G"
- apply (unfold Units_def)
- apply (clarsimp)
- apply (rule_tac x = "xaa \<otimes> xa" in bexI)
- apply auto
- apply (subst m_assoc)
- apply auto
- apply (subst (2) m_assoc [symmetric])
- apply auto
- apply (subst m_assoc)
- apply auto
- apply (subst (2) m_assoc [symmetric])
- apply auto
-done
-
-*)
-
-lemma (in monoid) units_group: "group(units_of G)"
- apply (unfold units_of_def)
- apply (rule groupI)
- apply auto
- apply (subst m_assoc)
- apply auto
- apply (rule_tac x = "inv x" in bexI)
- apply auto
-done
-
-lemma (in comm_monoid) units_comm_group: "comm_group(units_of G)"
- apply (rule group.group_comm_groupI)
- apply (rule units_group)
- apply (insert prems)
- apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
- apply auto;
-done;
-
-lemma units_of_carrier: "carrier (units_of G) = Units G"
- by (unfold units_of_def, auto)
-
-lemma units_of_mult: "mult(units_of G) = mult G"
- by (unfold units_of_def, auto)
-
-lemma units_of_one: "one(units_of G) = one G"
- by (unfold units_of_def, auto)
-
-lemma (in monoid) units_of_inv: "x : Units G ==>
- m_inv (units_of G) x = m_inv G x"
- apply (rule sym)
- apply (subst m_inv_def)
- apply (rule the1_equality)
- apply (rule ex_ex1I)
- apply (subst (asm) Units_def)
- apply auto
- apply (erule inv_unique)
- apply auto
- apply (rule Units_closed)
- apply (simp_all only: units_of_carrier [symmetric])
- apply (insert units_group)
- apply auto
- apply (subst units_of_mult [symmetric])
- apply (subst units_of_one [symmetric])
- apply (erule group.r_inv, assumption)
- apply (subst units_of_mult [symmetric])
- apply (subst units_of_one [symmetric])
- apply (erule group.l_inv, assumption)
-done
-
-lemma (in group) inj_on_const_mult: "a: (carrier G) ==>
- inj_on (%x. a \<otimes> x) (carrier G)"
- by (unfold inj_on_def, auto)
-
-lemma (in group) surj_const_mult: "a : (carrier G) ==>
- (%x. a \<otimes> x) ` (carrier G) = (carrier G)"
- apply (auto simp add: image_def)
- apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
- apply auto
-(* auto should get this. I suppose we need "comm_monoid_simprules"
- for mult_ac rewriting. *)
- apply (subst m_assoc [symmetric])
- apply auto
-done
-
-lemma (in group) l_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
- (x \<otimes> a = x) = (a = one G)"
- apply auto
- apply (subst l_cancel [symmetric])
- prefer 4
- apply (erule ssubst)
- apply auto
-done
-
-lemma (in group) r_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
- (a \<otimes> x = x) = (a = one G)"
- apply auto
- apply (subst r_cancel [symmetric])
- prefer 4
- apply (erule ssubst)
- apply auto
-done
-
-(* Is there a better way to do this? *)
-
-lemma (in group) l_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
- (x = x \<otimes> a) = (a = one G)"
- by (subst eq_commute, simp)
-
-lemma (in group) r_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
- (x = a \<otimes> x) = (a = one G)"
- by (subst eq_commute, simp)
-
-(* This should be generalized to arbitrary groups, not just commutative
- ones, using Lagrange's theorem. *)
-
-lemma (in comm_group) power_order_eq_one:
- assumes fin [simp]: "finite (carrier G)"
- and a [simp]: "a : carrier G"
- shows "a (^) card(carrier G) = one G"
-proof -
- have "(\<Otimes>x:carrier G. x) = (\<Otimes>x:carrier G. a \<otimes> x)"
- by (subst (2) finprod_reindex [symmetric],
- auto simp add: Pi_def inj_on_const_mult surj_const_mult)
- also have "\<dots> = (\<Otimes>x:carrier G. a) \<otimes> (\<Otimes>x:carrier G. x)"
- by (auto simp add: finprod_multf Pi_def)
- also have "(\<Otimes>x:carrier G. a) = a (^) card(carrier G)"
- by (auto simp add: finprod_const)
- finally show ?thesis
-(* uses the preceeding lemma *)
- by auto
-qed
-
-
-(* Miscellaneous *)
-
-lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> ~= \<one>\<^bsub>R\<^esub> \<Longrightarrow> ALL x : carrier R - {\<zero>\<^bsub>R\<^esub>}.
- x : Units R \<Longrightarrow> field R"
- apply (unfold_locales)
- apply (insert prems, auto)
- apply (rule trans)
- apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
- apply assumption
- apply (subst m_assoc)
- apply (auto simp add: Units_r_inv)
- apply (unfold Units_def)
- apply auto
-done
-
-lemma (in monoid) inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
- x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
- apply (subgoal_tac "x : Units G")
- apply (subgoal_tac "y = inv x \<otimes> \<one>")
- apply simp
- apply (erule subst)
- apply (subst m_assoc [symmetric])
- apply auto
- apply (unfold Units_def)
- apply auto
-done
-
-lemma (in comm_monoid) comm_inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
- x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
- apply (rule inv_char)
- apply auto
- apply (subst m_comm, auto)
-done
-
-lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"
- apply (rule inv_char)
- apply (auto simp add: l_minus r_minus)
-done
-
-lemma (in monoid) inv_eq_imp_eq: "x : Units G \<Longrightarrow> y : Units G \<Longrightarrow>
- inv x = inv y \<Longrightarrow> x = y"
- apply (subgoal_tac "inv(inv x) = inv(inv y)")
- apply (subst (asm) Units_inv_inv)+
- apply auto
-done
-
-lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> : Units R"
- apply (unfold Units_def)
- apply auto
- apply (rule_tac x = "\<ominus> \<one>" in bexI)
- apply auto
- apply (simp add: l_minus r_minus)
-done
-
-lemma (in monoid) inv_one [simp]: "inv \<one> = \<one>"
- apply (rule inv_char)
- apply auto
-done
-
-lemma (in ring) inv_eq_neg_one_eq: "x : Units R \<Longrightarrow> (inv x = \<ominus> \<one>) = (x = \<ominus> \<one>)"
- apply auto
- apply (subst Units_inv_inv [symmetric])
- apply auto
-done
-
-lemma (in monoid) inv_eq_one_eq: "x : Units G \<Longrightarrow> (inv x = \<one>) = (x = \<one>)"
- apply auto
- apply (subst Units_inv_inv [symmetric])
- apply auto
-done
-
-
-(* This goes in FiniteProduct *)
-
-lemma (in comm_monoid) finprod_UN_disjoint:
- "finite I \<Longrightarrow> (ALL i:I. finite (A i)) \<longrightarrow> (ALL i:I. ALL j:I. i ~= j \<longrightarrow>
- (A i) Int (A j) = {}) \<longrightarrow>
- (ALL i:I. ALL x: (A i). g x : carrier G) \<longrightarrow>
- finprod G g (UNION I A) = finprod G (%i. finprod G g (A i)) I"
- apply (induct set: finite)
- apply force
- apply clarsimp
- apply (subst finprod_Un_disjoint)
- apply blast
- apply (erule finite_UN_I)
- apply blast
- apply (fastsimp)
- apply (auto intro!: funcsetI finprod_closed)
-done
-
-lemma (in comm_monoid) finprod_Union_disjoint:
- "[| finite C; (ALL A:C. finite A & (ALL x:A. f x : carrier G));
- (ALL A:C. ALL B:C. A ~= B --> A Int B = {}) |]
- ==> finprod G f (Union C) = finprod G (finprod G f) C"
- apply (frule finprod_UN_disjoint [of C id f])
- apply (unfold Union_def id_def, auto)
-done
-
-lemma (in comm_monoid) finprod_one [rule_format]:
- "finite A \<Longrightarrow> (ALL x:A. f x = \<one>) \<longrightarrow>
- finprod G f A = \<one>"
-by (induct set: finite) auto
-
-
-(* need better simplification rules for rings *)
-(* the next one holds more generally for abelian groups *)
-
-lemma (in cring) sum_zero_eq_neg:
- "x : carrier R \<Longrightarrow> y : carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
- apply (subgoal_tac "\<ominus> y = \<zero> \<oplus> \<ominus> y")
- apply (erule ssubst)back
- apply (erule subst)
- apply (simp add: ring_simprules)+
-done
-
-(* there's a name conflict -- maybe "domain" should be
- "integral_domain" *)
-
-lemma (in Ring.domain) square_eq_one:
- fixes x
- assumes [simp]: "x : carrier R" and
- "x \<otimes> x = \<one>"
- shows "x = \<one> | x = \<ominus>\<one>"
-proof -
- have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
- by (simp add: ring_simprules)
- also with `x \<otimes> x = \<one>` have "\<dots> = \<zero>"
- by (simp add: ring_simprules)
- finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
- hence "(x \<oplus> \<one>) = \<zero> | (x \<oplus> \<ominus> \<one>) = \<zero>"
- by (intro integral, auto)
- thus ?thesis
- apply auto
- apply (erule notE)
- apply (rule sum_zero_eq_neg)
- apply auto
- apply (subgoal_tac "x = \<ominus> (\<ominus> \<one>)")
- apply (simp add: ring_simprules)
- apply (rule sum_zero_eq_neg)
- apply auto
- done
-qed
-
-lemma (in Ring.domain) inv_eq_self: "x : Units R \<Longrightarrow>
- x = inv x \<Longrightarrow> x = \<one> | x = \<ominus> \<one>"
- apply (rule square_eq_one)
- apply auto
- apply (erule ssubst)back
- apply (erule Units_r_inv)
-done
-
-
-(*
- The following translates theorems about groups to the facts about
- the units of a ring. (The list should be expanded as more things are
- needed.)
-*)
-
-lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow>
- finite (Units R)"
- by (rule finite_subset, auto)
-
-(* this belongs with MiscAlgebra.thy *)
-lemma (in monoid) units_of_pow:
- "x : Units G \<Longrightarrow> x (^)\<^bsub>units_of G\<^esub> (n::nat) = x (^)\<^bsub>G\<^esub> n"
- apply (induct n)
- apply (auto simp add: units_group group.is_monoid
- monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult
- One_nat_def)
-done
-
-lemma (in cring) units_power_order_eq_one: "finite (Units R) \<Longrightarrow> a : Units R
- \<Longrightarrow> a (^) card(Units R) = \<one>"
- apply (subst units_of_carrier [symmetric])
- apply (subst units_of_one [symmetric])
- apply (subst units_of_pow [symmetric])
- apply assumption
- apply (rule comm_group.power_order_eq_one)
- apply (rule units_comm_group)
- apply (unfold units_of_def, auto)
-done
-
-end
--- a/src/HOL/NewNumberTheory/ROOT.ML Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thys ["Fib","Residues"];
--- a/src/HOL/NewNumberTheory/Residues.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,466 +0,0 @@
-(* Title: HOL/Library/Residues.thy
- ID:
- Author: Jeremy Avigad
-
- An algebraic treatment of residue rings, and resulting proofs of
- Euler's theorem and Wilson's theorem.
-*)
-
-header {* Residue rings *}
-
-theory Residues
-imports
- UniqueFactorization
- Binomial
- MiscAlgebra
-begin
-
-
-(*
-
- A locale for residue rings
-
-*)
-
-constdefs
- residue_ring :: "int => int ring"
- "residue_ring m == (|
- carrier = {0..m - 1},
- mult = (%x y. (x * y) mod m),
- one = 1,
- zero = 0,
- add = (%x y. (x + y) mod m) |)"
-
-locale residues =
- fixes m :: int and R (structure)
- assumes m_gt_one: "m > 1"
- defines "R == residue_ring m"
-
-context residues begin
-
-lemma abelian_group: "abelian_group R"
- apply (insert m_gt_one)
- apply (rule abelian_groupI)
- apply (unfold R_def residue_ring_def)
- apply (auto simp add: mod_pos_pos_trivial mod_add_right_eq [symmetric]
- add_ac)
- apply (case_tac "x = 0")
- apply force
- apply (subgoal_tac "(x + (m - x)) mod m = 0")
- apply (erule bexI)
- apply auto
-done
-
-lemma comm_monoid: "comm_monoid R"
- apply (insert m_gt_one)
- apply (unfold R_def residue_ring_def)
- apply (rule comm_monoidI)
- apply auto
- apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m")
- apply (erule ssubst)
- apply (subst zmod_zmult1_eq [symmetric])+
- apply (simp_all only: mult_ac)
-done
-
-lemma cring: "cring R"
- apply (rule cringI)
- apply (rule abelian_group)
- apply (rule comm_monoid)
- apply (unfold R_def residue_ring_def, auto)
- apply (subst mod_add_eq [symmetric])
- apply (subst mult_commute)
- apply (subst zmod_zmult1_eq [symmetric])
- apply (simp add: ring_simps)
-done
-
-end
-
-sublocale residues < cring
- by (rule cring)
-
-
-context residues begin
-
-(* These lemmas translate back and forth between internal and
- external concepts *)
-
-lemma res_carrier_eq: "carrier R = {0..m - 1}"
- by (unfold R_def residue_ring_def, auto)
-
-lemma res_add_eq: "x \<oplus> y = (x + y) mod m"
- by (unfold R_def residue_ring_def, auto)
-
-lemma res_mult_eq: "x \<otimes> y = (x * y) mod m"
- by (unfold R_def residue_ring_def, auto)
-
-lemma res_zero_eq: "\<zero> = 0"
- by (unfold R_def residue_ring_def, auto)
-
-lemma res_one_eq: "\<one> = 1"
- by (unfold R_def residue_ring_def units_of_def residue_ring_def, auto)
-
-lemma res_units_eq: "Units R = { x. 0 < x & x < m & coprime x m}"
- apply (insert m_gt_one)
- apply (unfold Units_def R_def residue_ring_def)
- apply auto
- apply (subgoal_tac "x ~= 0")
- apply auto
- apply (rule invertible_coprime_int)
- apply (subgoal_tac "x ~= 0")
- apply auto
- apply (subst (asm) coprime_iff_invertible'_int)
- apply (rule m_gt_one)
- apply (auto simp add: cong_int_def mult_commute)
-done
-
-lemma res_neg_eq: "\<ominus> x = (- x) mod m"
- apply (insert m_gt_one)
- apply (unfold R_def a_inv_def m_inv_def residue_ring_def)
- apply auto
- apply (rule the_equality)
- apply auto
- apply (subst mod_add_right_eq [symmetric])
- apply auto
- apply (subst mod_add_left_eq [symmetric])
- apply auto
- apply (subgoal_tac "y mod m = - x mod m")
- apply simp
- apply (subst zmod_eq_dvd_iff)
- apply auto
-done
-
-lemma finite [iff]: "finite(carrier R)"
- by (subst res_carrier_eq, auto)
-
-lemma finite_Units [iff]: "finite(Units R)"
- by (subst res_units_eq, auto)
-
-(* The function a -> a mod m maps the integers to the
- residue classes. The following lemmas show that this mapping
- respects addition and multiplication on the integers. *)
-
-lemma mod_in_carrier [iff]: "a mod m : carrier R"
- apply (unfold res_carrier_eq)
- apply (insert m_gt_one, auto)
-done
-
-lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m"
- by (unfold R_def residue_ring_def, auto, arith)
-
-lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m"
- apply (unfold R_def residue_ring_def, auto)
- apply (subst zmod_zmult1_eq [symmetric])
- apply (subst mult_commute)
- apply (subst zmod_zmult1_eq [symmetric])
- apply (subst mult_commute)
- apply auto
-done
-
-lemma zero_cong: "\<zero> = 0"
- apply (unfold R_def residue_ring_def, auto)
-done
-
-lemma one_cong: "\<one> = 1 mod m"
- apply (insert m_gt_one)
- apply (unfold R_def residue_ring_def, auto)
-done
-
-(* revise algebra library to use 1? *)
-lemma pow_cong: "(x mod m) (^) n = x^n mod m"
- apply (insert m_gt_one)
- apply (induct n)
- apply (auto simp add: nat_pow_def one_cong One_nat_def)
- apply (subst mult_commute)
- apply (rule mult_cong)
-done
-
-lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m"
- apply (rule sym)
- apply (rule sum_zero_eq_neg)
- apply auto
- apply (subst add_cong)
- apply (subst zero_cong)
- apply auto
-done
-
-lemma (in residues) prod_cong:
- "finite A \<Longrightarrow> (\<Otimes> i:A. (f i) mod m) = (PROD i:A. f i) mod m"
- apply (induct set: finite)
- apply (auto simp: one_cong mult_cong)
-done
-
-lemma (in residues) sum_cong:
- "finite A \<Longrightarrow> (\<Oplus> i:A. (f i) mod m) = (SUM i: A. f i) mod m"
- apply (induct set: finite)
- apply (auto simp: zero_cong add_cong)
-done
-
-lemma mod_in_res_units [simp]: "1 < m \<Longrightarrow> coprime a m \<Longrightarrow>
- a mod m : Units R"
- apply (subst res_units_eq, auto)
- apply (insert pos_mod_sign [of m a])
- apply (subgoal_tac "a mod m ~= 0")
- apply arith
- apply auto
- apply (subst (asm) gcd_red_int)
- apply (subst gcd_commute_int, assumption)
-done
-
-lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))"
- unfolding cong_int_def by auto
-
-(* Simplifying with these will translate a ring equation in R to a
- congruence. *)
-
-lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong
- prod_cong sum_cong neg_cong res_eq_to_cong
-
-(* Other useful facts about the residue ring *)
-
-lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2"
- apply (simp add: res_one_eq res_neg_eq)
- apply (insert m_gt_one)
- apply (subgoal_tac "~(m > 2)")
- apply arith
- apply (rule notI)
- apply (subgoal_tac "-1 mod m = m - 1")
- apply force
- apply (subst mod_add_self2 [symmetric])
- apply (subst mod_pos_pos_trivial)
- apply auto
-done
-
-end
-
-
-(* prime residues *)
-
-locale residues_prime =
- fixes p :: int and R (structure)
- assumes p_prime [intro]: "prime p"
- defines "R == residue_ring p"
-
-sublocale residues_prime < residues p
- apply (unfold R_def residues_def)
- using p_prime apply auto
-done
-
-context residues_prime begin
-
-lemma is_field: "field R"
- apply (rule cring.field_intro2)
- apply (rule cring)
- apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq
- res_units_eq)
- apply (rule classical)
- apply (erule notE)
- apply (subst gcd_commute_int)
- apply (rule prime_imp_coprime_int)
- apply (rule p_prime)
- apply (rule notI)
- apply (frule zdvd_imp_le)
- apply auto
-done
-
-lemma res_prime_units_eq: "Units R = {1..p - 1}"
- apply (subst res_units_eq)
- apply auto
- apply (subst gcd_commute_int)
- apply (rule prime_imp_coprime_int)
- apply (rule p_prime)
- apply (rule zdvd_not_zless)
- apply auto
-done
-
-end
-
-sublocale residues_prime < field
- by (rule is_field)
-
-
-(*
- Test cases: Euler's theorem and Wilson's theorem.
-*)
-
-
-subsection{* Euler's theorem *}
-
-(* the definition of the phi function *)
-
-constdefs
- phi :: "int => nat"
- "phi m == card({ x. 0 < x & x < m & gcd x m = 1})"
-
-lemma phi_zero [simp]: "phi 0 = 0"
- apply (subst phi_def)
-(* Auto hangs here. Once again, where is the simplification rule
- 1 == Suc 0 coming from? *)
- apply (auto simp add: card_eq_0_iff)
-(* Add card_eq_0_iff as a simp rule? delete card_empty_imp? *)
-done
-
-lemma phi_one [simp]: "phi 1 = 0"
- apply (auto simp add: phi_def card_eq_0_iff)
-done
-
-lemma (in residues) phi_eq: "phi m = card(Units R)"
- by (simp add: phi_def res_units_eq)
-
-lemma (in residues) euler_theorem1:
- assumes a: "gcd a m = 1"
- shows "[a^phi m = 1] (mod m)"
-proof -
- from a m_gt_one have [simp]: "a mod m : Units R"
- by (intro mod_in_res_units)
- from phi_eq have "(a mod m) (^) (phi m) = (a mod m) (^) (card (Units R))"
- by simp
- also have "\<dots> = \<one>"
- by (intro units_power_order_eq_one, auto)
- finally show ?thesis
- by (simp add: res_to_cong_simps)
-qed
-
-(* In fact, there is a two line proof!
-
-lemma (in residues) euler_theorem1:
- assumes a: "gcd a m = 1"
- shows "[a^phi m = 1] (mod m)"
-proof -
- have "(a mod m) (^) (phi m) = \<one>"
- by (simp add: phi_eq units_power_order_eq_one a m_gt_one)
- thus ?thesis
- by (simp add: res_to_cong_simps)
-qed
-
-*)
-
-(* outside the locale, we can relax the restriction m > 1 *)
-
-lemma euler_theorem:
- assumes "m >= 0" and "gcd a m = 1"
- shows "[a^phi m = 1] (mod m)"
-proof (cases)
- assume "m = 0 | m = 1"
- thus ?thesis by auto
-next
- assume "~(m = 0 | m = 1)"
- with prems show ?thesis
- by (intro residues.euler_theorem1, unfold residues_def, auto)
-qed
-
-lemma (in residues_prime) phi_prime: "phi p = (nat p - 1)"
- apply (subst phi_eq)
- apply (subst res_prime_units_eq)
- apply auto
-done
-
-lemma phi_prime: "prime p \<Longrightarrow> phi p = (nat p - 1)"
- apply (rule residues_prime.phi_prime)
- apply (erule residues_prime.intro)
-done
-
-lemma fermat_theorem:
- assumes "prime p" and "~ (p dvd a)"
- shows "[a^(nat p - 1) = 1] (mod p)"
-proof -
- from prems have "[a^phi p = 1] (mod p)"
- apply (intro euler_theorem)
- (* auto should get this next part. matching across
- substitutions is needed. *)
- apply (frule prime_gt_1_int, arith)
- apply (subst gcd_commute_int, erule prime_imp_coprime_int, assumption)
- done
- also have "phi p = nat p - 1"
- by (rule phi_prime, rule prems)
- finally show ?thesis .
-qed
-
-
-subsection {* Wilson's theorem *}
-
-lemma (in field) inv_pair_lemma: "x : Units R \<Longrightarrow> y : Units R \<Longrightarrow>
- {x, inv x} ~= {y, inv y} \<Longrightarrow> {x, inv x} Int {y, inv y} = {}"
- apply auto
- apply (erule notE)
- apply (erule inv_eq_imp_eq)
- apply auto
- apply (erule notE)
- apply (erule inv_eq_imp_eq)
- apply auto
-done
-
-lemma (in residues_prime) wilson_theorem1:
- assumes a: "p > 2"
- shows "[fact (p - 1) = - 1] (mod p)"
-proof -
- let ?InversePairs = "{ {x, inv x} | x. x : Units R - {\<one>, \<ominus> \<one>}}"
- have UR: "Units R = {\<one>, \<ominus> \<one>} Un (Union ?InversePairs)"
- by auto
- have "(\<Otimes>i: Units R. i) =
- (\<Otimes>i: {\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i: Union ?InversePairs. i)"
- apply (subst UR)
- apply (subst finprod_Un_disjoint)
- apply (auto intro:funcsetI)
- apply (drule sym, subst (asm) inv_eq_one_eq)
- apply auto
- apply (drule sym, subst (asm) inv_eq_neg_one_eq)
- apply auto
- done
- also have "(\<Otimes>i: {\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
- apply (subst finprod_insert)
- apply auto
- apply (frule one_eq_neg_one)
- apply (insert a, force)
- done
- also have "(\<Otimes>i:(Union ?InversePairs). i) =
- (\<Otimes> A: ?InversePairs. (\<Otimes> y:A. y))"
- apply (subst finprod_Union_disjoint)
- apply force
- apply force
- apply clarify
- apply (rule inv_pair_lemma)
- apply auto
- done
- also have "\<dots> = \<one>"
- apply (rule finprod_one)
- apply auto
- apply (subst finprod_insert)
- apply auto
- apply (frule inv_eq_self)
- apply (auto)
- done
- finally have "(\<Otimes>i: Units R. i) = \<ominus> \<one>"
- by simp
- also have "(\<Otimes>i: Units R. i) = (\<Otimes>i: Units R. i mod p)"
- apply (rule finprod_cong')
- apply (auto)
- apply (subst (asm) res_prime_units_eq)
- apply auto
- done
- also have "\<dots> = (PROD i: Units R. i) mod p"
- apply (rule prod_cong)
- apply auto
- done
- also have "\<dots> = fact (p - 1) mod p"
- apply (subst fact_altdef_int)
- apply (insert prems, force)
- apply (subst res_prime_units_eq, rule refl)
- done
- finally have "fact (p - 1) mod p = \<ominus> \<one>".
- thus ?thesis
- by (simp add: res_to_cong_simps)
-qed
-
-lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)"
- apply (frule prime_gt_1_int)
- apply (case_tac "p = 2")
- apply (subst fact_altdef_int, simp)
- apply (subst cong_int_def)
- apply simp
- apply (rule residues_prime.wilson_theorem1)
- apply (rule residues_prime.intro)
- apply auto
-done
-
-
-end
--- a/src/HOL/NewNumberTheory/UniqueFactorization.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,967 +0,0 @@
-(* Title: UniqueFactorization.thy
- ID:
- Author: Jeremy Avigad
-
-
- Unique factorization for the natural numbers and the integers.
-
- Note: there were previous Isabelle formalizations of unique
- factorization due to Thomas Marthedal Rasmussen, and, building on
- that, by Jeremy Avigad and David Gray.
-*)
-
-header {* UniqueFactorization *}
-
-theory UniqueFactorization
-imports Cong Multiset
-begin
-
-(* inherited from Multiset *)
-declare One_nat_def [simp del]
-
-(* As a simp or intro rule,
-
- prime p \<Longrightarrow> p > 0
-
- wreaks havoc here. When the premise includes ALL x :# M. prime x, it
- leads to the backchaining
-
- x > 0
- prime x
- x :# M which is, unfortunately,
- count M x > 0
-*)
-
-
-(* useful facts *)
-
-lemma setsum_Un2: "finite (A Un B) \<Longrightarrow>
- setsum f (A Un B) = setsum f (A - B) + setsum f (B - A) +
- setsum f (A Int B)"
- apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
- apply (erule ssubst)
- apply (subst setsum_Un_disjoint)
- apply auto
- apply (subst setsum_Un_disjoint)
- apply auto
-done
-
-lemma setprod_Un2: "finite (A Un B) \<Longrightarrow>
- setprod f (A Un B) = setprod f (A - B) * setprod f (B - A) *
- setprod f (A Int B)"
- apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
- apply (erule ssubst)
- apply (subst setprod_Un_disjoint)
- apply auto
- apply (subst setprod_Un_disjoint)
- apply auto
-done
-
-(* Should this go in Multiset.thy? *)
-(* TN: No longer an intro-rule; needed only once and might get in the way *)
-lemma multiset_eqI: "[| !!x. count M x = count N x |] ==> M = N"
- by (subst multiset_eq_conv_count_eq, blast)
-
-(* Here is a version of set product for multisets. Is it worth moving
- to multiset.thy? If so, one should similarly define msetsum for abelian
- semirings, using of_nat. Also, is it worth developing bounded quantifiers
- "ALL i :# M. P i"?
-*)
-
-constdefs
- msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b"
- "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)"
-
-syntax
- "_msetprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"
- ("(3PROD _:#_. _)" [0, 51, 10] 10)
-
-translations
- "PROD i :# A. b" == "msetprod (%i. b) A"
-
-lemma msetprod_Un: "msetprod f (A+B) = msetprod f A * msetprod f B"
- apply (simp add: msetprod_def power_add)
- apply (subst setprod_Un2)
- apply auto
- apply (subgoal_tac
- "(PROD x:set_of A - set_of B. f x ^ count A x * f x ^ count B x) =
- (PROD x:set_of A - set_of B. f x ^ count A x)")
- apply (erule ssubst)
- apply (subgoal_tac
- "(PROD x:set_of B - set_of A. f x ^ count A x * f x ^ count B x) =
- (PROD x:set_of B - set_of A. f x ^ count B x)")
- apply (erule ssubst)
- apply (subgoal_tac "(PROD x:set_of A. f x ^ count A x) =
- (PROD x:set_of A - set_of B. f x ^ count A x) *
- (PROD x:set_of A Int set_of B. f x ^ count A x)")
- apply (erule ssubst)
- apply (subgoal_tac "(PROD x:set_of B. f x ^ count B x) =
- (PROD x:set_of B - set_of A. f x ^ count B x) *
- (PROD x:set_of A Int set_of B. f x ^ count B x)")
- apply (erule ssubst)
- apply (subst setprod_timesf)
- apply (force simp add: mult_ac)
- apply (subst setprod_Un_disjoint [symmetric])
- apply (auto intro: setprod_cong)
- apply (subst setprod_Un_disjoint [symmetric])
- apply (auto intro: setprod_cong)
-done
-
-
-subsection {* unique factorization: multiset version *}
-
-lemma multiset_prime_factorization_exists [rule_format]: "n > 0 -->
- (EX M. (ALL (p::nat) : set_of M. prime p) & n = (PROD i :# M. i))"
-proof (rule nat_less_induct, clarify)
- fix n :: nat
- assume ih: "ALL m < n. 0 < m --> (EX M. (ALL p : set_of M. prime p) & m =
- (PROD i :# M. i))"
- assume "(n::nat) > 0"
- then have "n = 1 | (n > 1 & prime n) | (n > 1 & ~ prime n)"
- by arith
- moreover
- {
- assume "n = 1"
- then have "(ALL p : set_of {#}. prime p) & n = (PROD i :# {#}. i)"
- by (auto simp add: msetprod_def)
- }
- moreover
- {
- assume "n > 1" and "prime n"
- then have "(ALL p : set_of {# n #}. prime p) & n = (PROD i :# {# n #}. i)"
- by (auto simp add: msetprod_def)
- }
- moreover
- {
- assume "n > 1" and "~ prime n"
- from prems not_prime_eq_prod_nat
- obtain m k where "n = m * k & 1 < m & m < n & 1 < k & k < n"
- by blast
- with ih obtain Q R where "(ALL p : set_of Q. prime p) & m = (PROD i:#Q. i)"
- and "(ALL p: set_of R. prime p) & k = (PROD i:#R. i)"
- by blast
- hence "(ALL p: set_of (Q + R). prime p) & n = (PROD i :# Q + R. i)"
- by (auto simp add: prems msetprod_Un set_of_union)
- then have "EX M. (ALL p : set_of M. prime p) & n = (PROD i :# M. i)"..
- }
- ultimately show "EX M. (ALL p : set_of M. prime p) & n = (PROD i::nat:#M. i)"
- by blast
-qed
-
-lemma multiset_prime_factorization_unique_aux:
- fixes a :: nat
- assumes "(ALL p : set_of M. prime p)" and
- "(ALL p : set_of N. prime p)" and
- "(PROD i :# M. i) dvd (PROD i:# N. i)"
- shows
- "count M a <= count N a"
-proof cases
- assume "a : set_of M"
- with prems have a: "prime a"
- by auto
- with prems have "a ^ count M a dvd (PROD i :# M. i)"
- by (auto intro: dvd_setprod simp add: msetprod_def)
- also have "... dvd (PROD i :# N. i)"
- by (rule prems)
- also have "... = (PROD i : (set_of N). i ^ (count N i))"
- by (simp add: msetprod_def)
- also have "... =
- a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))"
- proof (cases)
- assume "a : set_of N"
- hence b: "set_of N = {a} Un (set_of N - {a})"
- by auto
- thus ?thesis
- by (subst (1) b, subst setprod_Un_disjoint, auto)
- next
- assume "a ~: set_of N"
- thus ?thesis
- by auto
- qed
- finally have "a ^ count M a dvd
- a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))".
- moreover have "coprime (a ^ count M a)
- (PROD i : (set_of N - {a}). i ^ (count N i))"
- apply (subst gcd_commute_nat)
- apply (rule setprod_coprime_nat)
- apply (rule primes_imp_powers_coprime_nat)
- apply (insert prems, auto)
- done
- ultimately have "a ^ count M a dvd a^(count N a)"
- by (elim coprime_dvd_mult_nat)
- with a show ?thesis
- by (intro power_dvd_imp_le, auto)
-next
- assume "a ~: set_of M"
- thus ?thesis by auto
-qed
-
-lemma multiset_prime_factorization_unique:
- assumes "(ALL (p::nat) : set_of M. prime p)" and
- "(ALL p : set_of N. prime p)" and
- "(PROD i :# M. i) = (PROD i:# N. i)"
- shows
- "M = N"
-proof -
- {
- fix a
- from prems have "count M a <= count N a"
- by (intro multiset_prime_factorization_unique_aux, auto)
- moreover from prems have "count N a <= count M a"
- by (intro multiset_prime_factorization_unique_aux, auto)
- ultimately have "count M a = count N a"
- by auto
- }
- thus ?thesis by (simp add:multiset_eq_conv_count_eq)
-qed
-
-constdefs
- multiset_prime_factorization :: "nat => nat multiset"
- "multiset_prime_factorization n ==
- if n > 0 then (THE M. ((ALL p : set_of M. prime p) &
- n = (PROD i :# M. i)))
- else {#}"
-
-lemma multiset_prime_factorization: "n > 0 ==>
- (ALL p : set_of (multiset_prime_factorization n). prime p) &
- n = (PROD i :# (multiset_prime_factorization n). i)"
- apply (unfold multiset_prime_factorization_def)
- apply clarsimp
- apply (frule multiset_prime_factorization_exists)
- apply clarify
- apply (rule theI)
- apply (insert multiset_prime_factorization_unique, blast)+
-done
-
-
-subsection {* Prime factors and multiplicity for nats and ints *}
-
-class unique_factorization =
-
-fixes
- multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" and
- prime_factors :: "'a \<Rightarrow> 'a set"
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: unique_factorization
-
-begin
-
-definition
- multiplicity_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
- "multiplicity_nat p n = count (multiset_prime_factorization n) p"
-
-definition
- prime_factors_nat :: "nat \<Rightarrow> nat set"
-where
- "prime_factors_nat n = set_of (multiset_prime_factorization n)"
-
-instance proof qed
-
-end
-
-(* definitions for the integers *)
-
-instantiation int :: unique_factorization
-
-begin
-
-definition
- multiplicity_int :: "int \<Rightarrow> int \<Rightarrow> nat"
-where
- "multiplicity_int p n = multiplicity (nat p) (nat n)"
-
-definition
- prime_factors_int :: "int \<Rightarrow> int set"
-where
- "prime_factors_int n = int ` (prime_factors (nat n))"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up transfer *}
-
-lemma transfer_nat_int_prime_factors:
- "prime_factors (nat n) = nat ` prime_factors n"
- unfolding prime_factors_int_def apply auto
- by (subst transfer_int_nat_set_return_embed, assumption)
-
-lemma transfer_nat_int_prime_factors_closure: "n >= 0 \<Longrightarrow>
- nat_set (prime_factors n)"
- by (auto simp add: nat_set_def prime_factors_int_def)
-
-lemma transfer_nat_int_multiplicity: "p >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
- multiplicity (nat p) (nat n) = multiplicity p n"
- by (auto simp add: multiplicity_int_def)
-
-declare TransferMorphism_nat_int[transfer add return:
- transfer_nat_int_prime_factors transfer_nat_int_prime_factors_closure
- transfer_nat_int_multiplicity]
-
-
-lemma transfer_int_nat_prime_factors:
- "prime_factors (int n) = int ` prime_factors n"
- unfolding prime_factors_int_def by auto
-
-lemma transfer_int_nat_prime_factors_closure: "is_nat n \<Longrightarrow>
- nat_set (prime_factors n)"
- by (simp only: transfer_nat_int_prime_factors_closure is_nat_def)
-
-lemma transfer_int_nat_multiplicity:
- "multiplicity (int p) (int n) = multiplicity p n"
- by (auto simp add: multiplicity_int_def)
-
-declare TransferMorphism_int_nat[transfer add return:
- transfer_int_nat_prime_factors transfer_int_nat_prime_factors_closure
- transfer_int_nat_multiplicity]
-
-
-subsection {* Properties of prime factors and multiplicity for nats and ints *}
-
-lemma prime_factors_ge_0_int [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
- by (unfold prime_factors_int_def, auto)
-
-lemma prime_factors_prime_nat [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
- apply (case_tac "n = 0")
- apply (simp add: prime_factors_nat_def multiset_prime_factorization_def)
- apply (auto simp add: prime_factors_nat_def multiset_prime_factorization)
-done
-
-lemma prime_factors_prime_int [intro]:
- assumes "n >= 0" and "p : prime_factors (n::int)"
- shows "prime p"
-
- apply (rule prime_factors_prime_nat [transferred, of n p])
- using prems apply auto
-done
-
-lemma prime_factors_gt_0_nat [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
- by (frule prime_factors_prime_nat, auto)
-
-lemma prime_factors_gt_0_int [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow>
- p > (0::int)"
- by (frule (1) prime_factors_prime_int, auto)
-
-lemma prime_factors_finite_nat [iff]: "finite (prime_factors (n::nat))"
- by (unfold prime_factors_nat_def, auto)
-
-lemma prime_factors_finite_int [iff]: "finite (prime_factors (n::int))"
- by (unfold prime_factors_int_def, auto)
-
-lemma prime_factors_altdef_nat: "prime_factors (n::nat) =
- {p. multiplicity p n > 0}"
- by (force simp add: prime_factors_nat_def multiplicity_nat_def)
-
-lemma prime_factors_altdef_int: "prime_factors (n::int) =
- {p. p >= 0 & multiplicity p n > 0}"
- apply (unfold prime_factors_int_def multiplicity_int_def)
- apply (subst prime_factors_altdef_nat)
- apply (auto simp add: image_def)
-done
-
-lemma prime_factorization_nat: "(n::nat) > 0 \<Longrightarrow>
- n = (PROD p : prime_factors n. p^(multiplicity p n))"
- by (frule multiset_prime_factorization,
- simp add: prime_factors_nat_def multiplicity_nat_def msetprod_def)
-
-thm prime_factorization_nat [transferred]
-
-lemma prime_factorization_int:
- assumes "(n::int) > 0"
- shows "n = (PROD p : prime_factors n. p^(multiplicity p n))"
-
- apply (rule prime_factorization_nat [transferred, of n])
- using prems apply auto
-done
-
-lemma neq_zero_eq_gt_zero_nat: "((x::nat) ~= 0) = (x > 0)"
- by auto
-
-lemma prime_factorization_unique_nat:
- "S = { (p::nat) . f p > 0} \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
- n = (PROD p : S. p^(f p)) \<Longrightarrow>
- S = prime_factors n & (ALL p. f p = multiplicity p n)"
- apply (subgoal_tac "multiset_prime_factorization n = Abs_multiset
- f")
- apply (unfold prime_factors_nat_def multiplicity_nat_def)
- apply (simp add: set_of_def count_def Abs_multiset_inverse multiset_def)
- apply (unfold multiset_prime_factorization_def)
- apply (subgoal_tac "n > 0")
- prefer 2
- apply force
- apply (subst if_P, assumption)
- apply (rule the1_equality)
- apply (rule ex_ex1I)
- apply (rule multiset_prime_factorization_exists, assumption)
- apply (rule multiset_prime_factorization_unique)
- apply force
- apply force
- apply force
- unfolding set_of_def count_def msetprod_def
- apply (subgoal_tac "f : multiset")
- apply (auto simp only: Abs_multiset_inverse)
- unfolding multiset_def apply force
-done
-
-lemma prime_factors_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
- finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
- prime_factors n = S"
- by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric],
- assumption+)
-
-lemma prime_factors_characterization'_nat:
- "finite {p. 0 < f (p::nat)} \<Longrightarrow>
- (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
- prime_factors (PROD p | 0 < f p . p ^ f p) = {p. 0 < f p}"
- apply (rule prime_factors_characterization_nat)
- apply auto
-done
-
-(* A minor glitch:*)
-
-thm prime_factors_characterization'_nat
- [where f = "%x. f (int (x::nat))",
- transferred direction: nat "op <= (0::int)", rule_format]
-
-(*
- Transfer isn't smart enough to know that the "0 < f p" should
- remain a comparison between nats. But the transfer still works.
-*)
-
-lemma primes_characterization'_int [rule_format]:
- "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
- (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
- prime_factors (PROD p | p >=0 & 0 < f p . p ^ f p) =
- {p. p >= 0 & 0 < f p}"
-
- apply (insert prime_factors_characterization'_nat
- [where f = "%x. f (int (x::nat))",
- transferred direction: nat "op <= (0::int)"])
- apply auto
-done
-
-lemma prime_factors_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
- finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
- prime_factors n = S"
- apply simp
- apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
- apply (simp only:)
- apply (subst primes_characterization'_int)
- apply auto
- apply (auto simp add: prime_ge_0_int)
-done
-
-lemma multiplicity_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow>
- finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
- multiplicity p n = f p"
- by (frule prime_factorization_unique_nat [THEN conjunct2, rule_format,
- symmetric], auto)
-
-lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
- (ALL p. 0 < f p \<longrightarrow> prime p) \<longrightarrow>
- multiplicity p (PROD p | 0 < f p . p ^ f p) = f p"
- apply (rule impI)+
- apply (rule multiplicity_characterization_nat)
- apply auto
-done
-
-lemma multiplicity_characterization'_int [rule_format]:
- "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
- (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> p >= 0 \<Longrightarrow>
- multiplicity p (PROD p | p >= 0 & 0 < f p . p ^ f p) = f p"
-
- apply (insert multiplicity_characterization'_nat
- [where f = "%x. f (int (x::nat))",
- transferred direction: nat "op <= (0::int)", rule_format])
- apply auto
-done
-
-lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
- finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
- p >= 0 \<Longrightarrow> multiplicity p n = f p"
- apply simp
- apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
- apply (simp only:)
- apply (subst multiplicity_characterization'_int)
- apply auto
- apply (auto simp add: prime_ge_0_int)
-done
-
-lemma multiplicity_zero_nat [simp]: "multiplicity (p::nat) 0 = 0"
- by (simp add: multiplicity_nat_def multiset_prime_factorization_def)
-
-lemma multiplicity_zero_int [simp]: "multiplicity (p::int) 0 = 0"
- by (simp add: multiplicity_int_def)
-
-lemma multiplicity_one_nat [simp]: "multiplicity p (1::nat) = 0"
- by (subst multiplicity_characterization_nat [where f = "%x. 0"], auto)
-
-lemma multiplicity_one_int [simp]: "multiplicity p (1::int) = 0"
- by (simp add: multiplicity_int_def)
-
-lemma multiplicity_prime_nat [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
- apply (subst multiplicity_characterization_nat
- [where f = "(%q. if q = p then 1 else 0)"])
- apply auto
- apply (case_tac "x = p")
- apply auto
-done
-
-lemma multiplicity_prime_int [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
- unfolding prime_int_def multiplicity_int_def by auto
-
-lemma multiplicity_prime_power_nat [simp]: "prime (p::nat) \<Longrightarrow>
- multiplicity p (p^n) = n"
- apply (case_tac "n = 0")
- apply auto
- apply (subst multiplicity_characterization_nat
- [where f = "(%q. if q = p then n else 0)"])
- apply auto
- apply (case_tac "x = p")
- apply auto
-done
-
-lemma multiplicity_prime_power_int [simp]: "prime (p::int) \<Longrightarrow>
- multiplicity p (p^n) = n"
- apply (frule prime_ge_0_int)
- apply (auto simp add: prime_int_def multiplicity_int_def nat_power_eq)
-done
-
-lemma multiplicity_nonprime_nat [simp]: "~ prime (p::nat) \<Longrightarrow>
- multiplicity p n = 0"
- apply (case_tac "n = 0")
- apply auto
- apply (frule multiset_prime_factorization)
- apply (auto simp add: set_of_def multiplicity_nat_def)
-done
-
-lemma multiplicity_nonprime_int [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
- by (unfold multiplicity_int_def prime_int_def, auto)
-
-lemma multiplicity_not_factor_nat [simp]:
- "p ~: prime_factors (n::nat) \<Longrightarrow> multiplicity p n = 0"
- by (subst (asm) prime_factors_altdef_nat, auto)
-
-lemma multiplicity_not_factor_int [simp]:
- "p >= 0 \<Longrightarrow> p ~: prime_factors (n::int) \<Longrightarrow> multiplicity p n = 0"
- by (subst (asm) prime_factors_altdef_int, auto)
-
-lemma multiplicity_product_aux_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
- (prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
- (ALL p. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
- apply (rule prime_factorization_unique_nat)
- apply (simp only: prime_factors_altdef_nat)
- apply auto
- apply (subst power_add)
- apply (subst setprod_timesf)
- apply (rule arg_cong2)back back
- apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors k Un
- (prime_factors l - prime_factors k)")
- apply (erule ssubst)
- apply (subst setprod_Un_disjoint)
- apply auto
- apply (subgoal_tac "(\<Prod>p\<in>prime_factors l - prime_factors k. p ^ multiplicity p k) =
- (\<Prod>p\<in>prime_factors l - prime_factors k. 1)")
- apply (erule ssubst)
- apply (simp add: setprod_1)
- apply (erule prime_factorization_nat)
- apply (rule setprod_cong, auto)
- apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un
- (prime_factors k - prime_factors l)")
- apply (erule ssubst)
- apply (subst setprod_Un_disjoint)
- apply auto
- apply (subgoal_tac "(\<Prod>p\<in>prime_factors k - prime_factors l. p ^ multiplicity p l) =
- (\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
- apply (erule ssubst)
- apply (simp add: setprod_1)
- apply (erule prime_factorization_nat)
- apply (rule setprod_cong, auto)
-done
-
-(* transfer doesn't have the same problem here with the right
- choice of rules. *)
-
-lemma multiplicity_product_aux_int:
- assumes "(k::int) > 0" and "l > 0"
- shows
- "(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
- (ALL p >= 0. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
-
- apply (rule multiplicity_product_aux_nat [transferred, of l k])
- using prems apply auto
-done
-
-lemma prime_factors_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
- prime_factors k Un prime_factors l"
- by (rule multiplicity_product_aux_nat [THEN conjunct1, symmetric])
-
-lemma prime_factors_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) =
- prime_factors k Un prime_factors l"
- by (rule multiplicity_product_aux_int [THEN conjunct1, symmetric])
-
-lemma multiplicity_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) =
- multiplicity p k + multiplicity p l"
- by (rule multiplicity_product_aux_nat [THEN conjunct2, rule_format,
- symmetric])
-
-lemma multiplicity_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow>
- multiplicity p (k * l) = multiplicity p k + multiplicity p l"
- by (rule multiplicity_product_aux_int [THEN conjunct2, rule_format,
- symmetric])
-
-lemma multiplicity_setprod_nat: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow>
- multiplicity (p::nat) (PROD x : S. f x) =
- (SUM x : S. multiplicity p (f x))"
- apply (induct set: finite)
- apply auto
- apply (subst multiplicity_product_nat)
- apply auto
-done
-
-(* Transfer is delicate here for two reasons: first, because there is
- an implicit quantifier over functions (f), and, second, because the
- product over the multiplicity should not be translated to an integer
- product.
-
- The way to handle the first is to use quantifier rules for functions.
- The way to handle the second is to turn off the offending rule.
-*)
-
-lemma transfer_nat_int_sum_prod_closure3:
- "(SUM x : A. int (f x)) >= 0"
- "(PROD x : A. int (f x)) >= 0"
- apply (rule setsum_nonneg, auto)
- apply (rule setprod_nonneg, auto)
-done
-
-declare TransferMorphism_nat_int[transfer
- add return: transfer_nat_int_sum_prod_closure3
- del: transfer_nat_int_sum_prod2 (1)]
-
-lemma multiplicity_setprod_int: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow>
- (ALL x : S. f x > 0) \<Longrightarrow>
- multiplicity (p::int) (PROD x : S. f x) =
- (SUM x : S. multiplicity p (f x))"
-
- apply (frule multiplicity_setprod_nat
- [where f = "%x. nat(int(nat(f x)))",
- transferred direction: nat "op <= (0::int)"])
- apply auto
- apply (subst (asm) setprod_cong)
- apply (rule refl)
- apply (rule if_P)
- apply auto
- apply (rule setsum_cong)
- apply auto
-done
-
-declare TransferMorphism_nat_int[transfer
- add return: transfer_nat_int_sum_prod2 (1)]
-
-lemma multiplicity_prod_prime_powers_nat:
- "finite S \<Longrightarrow> (ALL p : S. prime (p::nat)) \<Longrightarrow>
- multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
- apply (subgoal_tac "(PROD p : S. p ^ f p) =
- (PROD p : S. p ^ (%x. if x : S then f x else 0) p)")
- apply (erule ssubst)
- apply (subst multiplicity_characterization_nat)
- prefer 5 apply (rule refl)
- apply (rule refl)
- apply auto
- apply (subst setprod_mono_one_right)
- apply assumption
- prefer 3
- apply (rule setprod_cong)
- apply (rule refl)
- apply auto
-done
-
-(* Here the issue with transfer is the implicit quantifier over S *)
-
-lemma multiplicity_prod_prime_powers_int:
- "(p::int) >= 0 \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
- multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
-
- apply (subgoal_tac "int ` nat ` S = S")
- apply (frule multiplicity_prod_prime_powers_nat [where f = "%x. f(int x)"
- and S = "nat ` S", transferred])
- apply auto
- apply (subst prime_int_def [symmetric])
- apply auto
- apply (subgoal_tac "xb >= 0")
- apply force
- apply (rule prime_ge_0_int)
- apply force
- apply (subst transfer_nat_int_set_return_embed)
- apply (unfold nat_set_def, auto)
-done
-
-lemma multiplicity_distinct_prime_power_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
- p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
- apply (subgoal_tac "q^n = setprod (%x. x^n) {q}")
- apply (erule ssubst)
- apply (subst multiplicity_prod_prime_powers_nat)
- apply auto
-done
-
-lemma multiplicity_distinct_prime_power_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
- p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
- apply (frule prime_ge_0_int [of q])
- apply (frule multiplicity_distinct_prime_power_nat [transferred leaving: n])
- prefer 4
- apply assumption
- apply auto
-done
-
-lemma dvd_multiplicity_nat:
- "(0::nat) < y \<Longrightarrow> x dvd y \<Longrightarrow> multiplicity p x <= multiplicity p y"
- apply (case_tac "x = 0")
- apply (auto simp add: dvd_def multiplicity_product_nat)
-done
-
-lemma dvd_multiplicity_int:
- "(0::int) < y \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> p >= 0 \<Longrightarrow>
- multiplicity p x <= multiplicity p y"
- apply (case_tac "x = 0")
- apply (auto simp add: dvd_def)
- apply (subgoal_tac "0 < k")
- apply (auto simp add: multiplicity_product_int)
- apply (erule zero_less_mult_pos)
- apply arith
-done
-
-lemma dvd_prime_factors_nat [intro]:
- "0 < (y::nat) \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
- apply (simp only: prime_factors_altdef_nat)
- apply auto
- apply (frule dvd_multiplicity_nat)
- apply auto
-(* It is a shame that auto and arith don't get this. *)
- apply (erule order_less_le_trans)back
- apply assumption
-done
-
-lemma dvd_prime_factors_int [intro]:
- "0 < (y::int) \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
- apply (auto simp add: prime_factors_altdef_int)
- apply (erule order_less_le_trans)
- apply (rule dvd_multiplicity_int)
- apply auto
-done
-
-lemma multiplicity_dvd_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
- ALL p. multiplicity p x <= multiplicity p y \<Longrightarrow>
- x dvd y"
- apply (subst prime_factorization_nat [of x], assumption)
- apply (subst prime_factorization_nat [of y], assumption)
- apply (rule setprod_dvd_setprod_subset2)
- apply force
- apply (subst prime_factors_altdef_nat)+
- apply auto
-(* Again, a shame that auto and arith don't get this. *)
- apply (drule_tac x = xa in spec, auto)
- apply (rule le_imp_power_dvd)
- apply blast
-done
-
-lemma multiplicity_dvd_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
- ALL p >= 0. multiplicity p x <= multiplicity p y \<Longrightarrow>
- x dvd y"
- apply (subst prime_factorization_int [of x], assumption)
- apply (subst prime_factorization_int [of y], assumption)
- apply (rule setprod_dvd_setprod_subset2)
- apply force
- apply (subst prime_factors_altdef_int)+
- apply auto
- apply (rule dvd_power_le)
- apply auto
- apply (drule_tac x = xa in spec)
- apply (erule impE)
- apply auto
-done
-
-lemma multiplicity_dvd'_nat: "(0::nat) < x \<Longrightarrow>
- \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
- apply (cases "y = 0")
- apply auto
- apply (rule multiplicity_dvd_nat, auto)
- apply (case_tac "prime p")
- apply auto
-done
-
-lemma multiplicity_dvd'_int: "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
- \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
- apply (cases "y = 0")
- apply auto
- apply (rule multiplicity_dvd_int, auto)
- apply (case_tac "prime p")
- apply auto
-done
-
-lemma dvd_multiplicity_eq_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
- (x dvd y) = (ALL p. multiplicity p x <= multiplicity p y)"
- by (auto intro: dvd_multiplicity_nat multiplicity_dvd_nat)
-
-lemma dvd_multiplicity_eq_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
- (x dvd y) = (ALL p >= 0. multiplicity p x <= multiplicity p y)"
- by (auto intro: dvd_multiplicity_int multiplicity_dvd_int)
-
-lemma prime_factors_altdef2_nat: "(n::nat) > 0 \<Longrightarrow>
- (p : prime_factors n) = (prime p & p dvd n)"
- apply (case_tac "prime p")
- apply auto
- apply (subst prime_factorization_nat [where n = n], assumption)
- apply (rule dvd_trans)
- apply (rule dvd_power [where x = p and n = "multiplicity p n"])
- apply (subst (asm) prime_factors_altdef_nat, force)
- apply (rule dvd_setprod)
- apply auto
- apply (subst prime_factors_altdef_nat)
- apply (subst (asm) dvd_multiplicity_eq_nat)
- apply auto
- apply (drule spec [where x = p])
- apply auto
-done
-
-lemma prime_factors_altdef2_int:
- assumes "(n::int) > 0"
- shows "(p : prime_factors n) = (prime p & p dvd n)"
-
- apply (case_tac "p >= 0")
- apply (rule prime_factors_altdef2_nat [transferred])
- using prems apply auto
- apply (auto simp add: prime_ge_0_int prime_factors_ge_0_int)
-done
-
-lemma multiplicity_eq_nat:
- fixes x and y::nat
- assumes [arith]: "x > 0" "y > 0" and
- mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
- shows "x = y"
-
- apply (rule dvd_anti_sym)
- apply (auto intro: multiplicity_dvd'_nat)
-done
-
-lemma multiplicity_eq_int:
- fixes x and y::int
- assumes [arith]: "x > 0" "y > 0" and
- mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
- shows "x = y"
-
- apply (rule dvd_anti_sym [transferred])
- apply (auto intro: multiplicity_dvd'_int)
-done
-
-
-subsection {* An application *}
-
-lemma gcd_eq_nat:
- assumes pos [arith]: "x > 0" "y > 0"
- shows "gcd (x::nat) y =
- (PROD p: prime_factors x Un prime_factors y.
- p ^ (min (multiplicity p x) (multiplicity p y)))"
-proof -
- def z == "(PROD p: prime_factors (x::nat) Un prime_factors y.
- p ^ (min (multiplicity p x) (multiplicity p y)))"
- have [arith]: "z > 0"
- unfolding z_def by (rule setprod_pos_nat, auto)
- have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
- min (multiplicity p x) (multiplicity p y)"
- unfolding z_def
- apply (subst multiplicity_prod_prime_powers_nat)
- apply (auto simp add: multiplicity_not_factor_nat)
- done
- have "z dvd x"
- by (intro multiplicity_dvd'_nat, auto simp add: aux)
- moreover have "z dvd y"
- by (intro multiplicity_dvd'_nat, auto simp add: aux)
- moreover have "ALL w. w dvd x & w dvd y \<longrightarrow> w dvd z"
- apply auto
- apply (case_tac "w = 0", auto)
- apply (erule multiplicity_dvd'_nat)
- apply (auto intro: dvd_multiplicity_nat simp add: aux)
- done
- ultimately have "z = gcd x y"
- by (subst gcd_unique_nat [symmetric], blast)
- thus ?thesis
- unfolding z_def by auto
-qed
-
-lemma lcm_eq_nat:
- assumes pos [arith]: "x > 0" "y > 0"
- shows "lcm (x::nat) y =
- (PROD p: prime_factors x Un prime_factors y.
- p ^ (max (multiplicity p x) (multiplicity p y)))"
-proof -
- def z == "(PROD p: prime_factors (x::nat) Un prime_factors y.
- p ^ (max (multiplicity p x) (multiplicity p y)))"
- have [arith]: "z > 0"
- unfolding z_def by (rule setprod_pos_nat, auto)
- have aux: "!!p. prime p \<Longrightarrow> multiplicity p z =
- max (multiplicity p x) (multiplicity p y)"
- unfolding z_def
- apply (subst multiplicity_prod_prime_powers_nat)
- apply (auto simp add: multiplicity_not_factor_nat)
- done
- have "x dvd z"
- by (intro multiplicity_dvd'_nat, auto simp add: aux)
- moreover have "y dvd z"
- by (intro multiplicity_dvd'_nat, auto simp add: aux)
- moreover have "ALL w. x dvd w & y dvd w \<longrightarrow> z dvd w"
- apply auto
- apply (case_tac "w = 0", auto)
- apply (rule multiplicity_dvd'_nat)
- apply (auto intro: dvd_multiplicity_nat simp add: aux)
- done
- ultimately have "z = lcm x y"
- by (subst lcm_unique_nat [symmetric], blast)
- thus ?thesis
- unfolding z_def by auto
-qed
-
-lemma multiplicity_gcd_nat:
- assumes [arith]: "x > 0" "y > 0"
- shows "multiplicity (p::nat) (gcd x y) =
- min (multiplicity p x) (multiplicity p y)"
-
- apply (subst gcd_eq_nat)
- apply auto
- apply (subst multiplicity_prod_prime_powers_nat)
- apply auto
-done
-
-lemma multiplicity_lcm_nat:
- assumes [arith]: "x > 0" "y > 0"
- shows "multiplicity (p::nat) (lcm x y) =
- max (multiplicity p x) (multiplicity p y)"
-
- apply (subst lcm_eq_nat)
- apply auto
- apply (subst multiplicity_prod_prime_powers_nat)
- apply auto
-done
-
-lemma gcd_lcm_distrib_nat: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
- apply (case_tac "x = 0 | y = 0 | z = 0")
- apply auto
- apply (rule multiplicity_eq_nat)
- apply (auto simp add: multiplicity_gcd_nat multiplicity_lcm_nat
- lcm_pos_nat)
-done
-
-lemma gcd_lcm_distrib_int: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
- apply (subst (1 2 3) gcd_abs_int)
- apply (subst lcm_abs_int)
- apply (subst (2) abs_of_nonneg)
- apply force
- apply (rule gcd_lcm_distrib_nat [transferred])
- apply auto
-done
-
-end
--- a/src/HOL/NumberTheory/BijectionRel.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,231 +0,0 @@
-(* Title: HOL/NumberTheory/BijectionRel.thy
- ID: $Id$
- Author: Thomas M. Rasmussen
- Copyright 2000 University of Cambridge
-*)
-
-header {* Bijections between sets *}
-
-theory BijectionRel imports Main begin
-
-text {*
- Inductive definitions of bijections between two different sets and
- between the same set. Theorem for relating the two definitions.
-
- \bigskip
-*}
-
-inductive_set
- bijR :: "('a => 'b => bool) => ('a set * 'b set) set"
- for P :: "'a => 'b => bool"
-where
- empty [simp]: "({}, {}) \<in> bijR P"
-| insert: "P a b ==> a \<notin> A ==> b \<notin> B ==> (A, B) \<in> bijR P
- ==> (insert a A, insert b B) \<in> bijR P"
-
-text {*
- Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
- (and similar for @{term A}).
-*}
-
-definition
- bijP :: "('a => 'a => bool) => 'a set => bool" where
- "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
-
-definition
- uniqP :: "('a => 'a => bool) => bool" where
- "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
-
-definition
- symP :: "('a => 'a => bool) => bool" where
- "symP P = (\<forall>a b. P a b = P b a)"
-
-inductive_set
- bijER :: "('a => 'a => bool) => 'a set set"
- for P :: "'a => 'a => bool"
-where
- empty [simp]: "{} \<in> bijER P"
-| insert1: "P a a ==> a \<notin> A ==> A \<in> bijER P ==> insert a A \<in> bijER P"
-| insert2: "P a b ==> a \<noteq> b ==> a \<notin> A ==> b \<notin> A ==> A \<in> bijER P
- ==> insert a (insert b A) \<in> bijER P"
-
-
-text {* \medskip @{term bijR} *}
-
-lemma fin_bijRl: "(A, B) \<in> bijR P ==> finite A"
- apply (erule bijR.induct)
- apply auto
- done
-
-lemma fin_bijRr: "(A, B) \<in> bijR P ==> finite B"
- apply (erule bijR.induct)
- apply auto
- done
-
-lemma aux_induct:
- assumes major: "finite F"
- and subs: "F \<subseteq> A"
- and cases: "P {}"
- "!!F a. F \<subseteq> A ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)"
- shows "P F"
- using major subs
- apply (induct set: finite)
- apply (blast intro: cases)+
- done
-
-
-lemma inj_func_bijR_aux1:
- "A \<subseteq> B ==> a \<notin> A ==> a \<in> B ==> inj_on f B ==> f a \<notin> f ` A"
- apply (unfold inj_on_def)
- apply auto
- done
-
-lemma inj_func_bijR_aux2:
- "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A ==> F <= A
- ==> (F, f ` F) \<in> bijR P"
- apply (rule_tac F = F and A = A in aux_induct)
- apply (rule finite_subset)
- apply auto
- apply (rule bijR.insert)
- apply (rule_tac [3] inj_func_bijR_aux1)
- apply auto
- done
-
-lemma inj_func_bijR:
- "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A
- ==> (A, f ` A) \<in> bijR P"
- apply (rule inj_func_bijR_aux2)
- apply auto
- done
-
-
-text {* \medskip @{term bijER} *}
-
-lemma fin_bijER: "A \<in> bijER P ==> finite A"
- apply (erule bijER.induct)
- apply auto
- done
-
-lemma aux1:
- "a \<notin> A ==> a \<notin> B ==> F \<subseteq> insert a A ==> F \<subseteq> insert a B ==> a \<in> F
- ==> \<exists>C. F = insert a C \<and> a \<notin> C \<and> C <= A \<and> C <= B"
- apply (rule_tac x = "F - {a}" in exI)
- apply auto
- done
-
-lemma aux2: "a \<noteq> b ==> a \<notin> A ==> b \<notin> B ==> a \<in> F ==> b \<in> F
- ==> F \<subseteq> insert a A ==> F \<subseteq> insert b B
- ==> \<exists>C. F = insert a (insert b C) \<and> a \<notin> C \<and> b \<notin> C \<and> C \<subseteq> A \<and> C \<subseteq> B"
- apply (rule_tac x = "F - {a, b}" in exI)
- apply auto
- done
-
-lemma aux_uniq: "uniqP P ==> P a b ==> P c d ==> (a = c) = (b = d)"
- apply (unfold uniqP_def)
- apply auto
- done
-
-lemma aux_sym: "symP P ==> P a b = P b a"
- apply (unfold symP_def)
- apply auto
- done
-
-lemma aux_in1:
- "uniqP P ==> b \<notin> C ==> P b b ==> bijP P (insert b C) ==> bijP P C"
- apply (unfold bijP_def)
- apply auto
- apply (subgoal_tac "b \<noteq> a")
- prefer 2
- apply clarify
- apply (simp add: aux_uniq)
- apply auto
- done
-
-lemma aux_in2:
- "symP P ==> uniqP P ==> a \<notin> C ==> b \<notin> C ==> a \<noteq> b ==> P a b
- ==> bijP P (insert a (insert b C)) ==> bijP P C"
- apply (unfold bijP_def)
- apply auto
- apply (subgoal_tac "aa \<noteq> a")
- prefer 2
- apply clarify
- apply (subgoal_tac "aa \<noteq> b")
- prefer 2
- apply clarify
- apply (simp add: aux_uniq)
- apply (subgoal_tac "ba \<noteq> a")
- apply auto
- apply (subgoal_tac "P a aa")
- prefer 2
- apply (simp add: aux_sym)
- apply (subgoal_tac "b = aa")
- apply (rule_tac [2] iffD1)
- apply (rule_tac [2] a = a and c = a and P = P in aux_uniq)
- apply auto
- done
-
-lemma aux_foo: "\<forall>a b. Q a \<and> P a b --> R b ==> P a b ==> Q a ==> R b"
- apply auto
- done
-
-lemma aux_bij: "bijP P F ==> symP P ==> P a b ==> (a \<in> F) = (b \<in> F)"
- apply (unfold bijP_def)
- apply (rule iffI)
- apply (erule_tac [!] aux_foo)
- apply simp_all
- apply (rule iffD2)
- apply (rule_tac P = P in aux_sym)
- apply simp_all
- done
-
-
-lemma aux_bijRER:
- "(A, B) \<in> bijR P ==> uniqP P ==> symP P
- ==> \<forall>F. bijP P F \<and> F \<subseteq> A \<and> F \<subseteq> B --> F \<in> bijER P"
- apply (erule bijR.induct)
- apply simp
- apply (case_tac "a = b")
- apply clarify
- apply (case_tac "b \<in> F")
- prefer 2
- apply (simp add: subset_insert)
- apply (cut_tac F = F and a = b and A = A and B = B in aux1)
- prefer 6
- apply clarify
- apply (rule bijER.insert1)
- apply simp_all
- apply (subgoal_tac "bijP P C")
- apply simp
- apply (rule aux_in1)
- apply simp_all
- apply clarify
- apply (case_tac "a \<in> F")
- apply (case_tac [!] "b \<in> F")
- apply (cut_tac F = F and a = a and b = b and A = A and B = B
- in aux2)
- apply (simp_all add: subset_insert)
- apply clarify
- apply (rule bijER.insert2)
- apply simp_all
- apply (subgoal_tac "bijP P C")
- apply simp
- apply (rule aux_in2)
- apply simp_all
- apply (subgoal_tac "b \<in> F")
- apply (rule_tac [2] iffD1)
- apply (rule_tac [2] a = a and F = F and P = P in aux_bij)
- apply (simp_all (no_asm_simp))
- apply (subgoal_tac [2] "a \<in> F")
- apply (rule_tac [3] iffD2)
- apply (rule_tac [3] b = b and F = F and P = P in aux_bij)
- apply auto
- done
-
-lemma bijR_bijER:
- "(A, A) \<in> bijR P ==>
- bijP P A ==> uniqP P ==> symP P ==> A \<in> bijER P"
- apply (cut_tac A = A and B = A and P = P in aux_bijRER)
- apply auto
- done
-
-end
--- a/src/HOL/NumberTheory/Chinese.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,259 +0,0 @@
-(* Title: HOL/NumberTheory/Chinese.thy
- ID: $Id$
- Author: Thomas M. Rasmussen
- Copyright 2000 University of Cambridge
-*)
-
-header {* The Chinese Remainder Theorem *}
-
-theory Chinese
-imports IntPrimes
-begin
-
-text {*
- The Chinese Remainder Theorem for an arbitrary finite number of
- equations. (The one-equation case is included in theory @{text
- IntPrimes}. Uses functions for indexing.\footnote{Maybe @{term
- funprod} and @{term funsum} should be based on general @{term fold}
- on indices?}
-*}
-
-
-subsection {* Definitions *}
-
-consts
- funprod :: "(nat => int) => nat => nat => int"
- funsum :: "(nat => int) => nat => nat => int"
-
-primrec
- "funprod f i 0 = f i"
- "funprod f i (Suc n) = f (Suc (i + n)) * funprod f i n"
-
-primrec
- "funsum f i 0 = f i"
- "funsum f i (Suc n) = f (Suc (i + n)) + funsum f i n"
-
-definition
- m_cond :: "nat => (nat => int) => bool" where
- "m_cond n mf =
- ((\<forall>i. i \<le> n --> 0 < mf i) \<and>
- (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i) (mf j) = 1))"
-
-definition
- km_cond :: "nat => (nat => int) => (nat => int) => bool" where
- "km_cond n kf mf = (\<forall>i. i \<le> n --> zgcd (kf i) (mf i) = 1)"
-
-definition
- lincong_sol ::
- "nat => (nat => int) => (nat => int) => (nat => int) => int => bool" where
- "lincong_sol n kf bf mf x = (\<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i))"
-
-definition
- mhf :: "(nat => int) => nat => nat => int" where
- "mhf mf n i =
- (if i = 0 then funprod mf (Suc 0) (n - Suc 0)
- else if i = n then funprod mf 0 (n - Suc 0)
- else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i))"
-
-definition
- xilin_sol ::
- "nat => nat => (nat => int) => (nat => int) => (nat => int) => int" where
- "xilin_sol i n kf bf mf =
- (if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then
- (SOME x. 0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
- else 0)"
-
-definition
- x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int" where
- "x_sol n kf bf mf = funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n"
-
-
-text {* \medskip @{term funprod} and @{term funsum} *}
-
-lemma funprod_pos: "(\<forall>i. i \<le> n --> 0 < mf i) ==> 0 < funprod mf 0 n"
- apply (induct n)
- apply auto
- apply (simp add: zero_less_mult_iff)
- done
-
-lemma funprod_zgcd [rule_format (no_asm)]:
- "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i) (mf m) = 1) -->
- zgcd (funprod mf k l) (mf m) = 1"
- apply (induct l)
- apply simp_all
- apply (rule impI)+
- apply (subst zgcd_zmult_cancel)
- apply auto
- done
-
-lemma funprod_zdvd [rule_format]:
- "k \<le> i --> i \<le> k + l --> mf i dvd funprod mf k l"
- apply (induct l)
- apply auto
- apply (subgoal_tac "i = Suc (k + l)")
- apply (simp_all (no_asm_simp))
- done
-
-lemma funsum_mod:
- "funsum f k l mod m = funsum (\<lambda>i. (f i) mod m) k l mod m"
- apply (induct l)
- apply auto
- apply (rule trans)
- apply (rule mod_add_eq)
- apply simp
- apply (rule mod_add_right_eq [symmetric])
- done
-
-lemma funsum_zero [rule_format (no_asm)]:
- "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = 0) --> (funsum f k l) = 0"
- apply (induct l)
- apply auto
- done
-
-lemma funsum_oneelem [rule_format (no_asm)]:
- "k \<le> j --> j \<le> k + l -->
- (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = 0) -->
- funsum f k l = f j"
- apply (induct l)
- prefer 2
- apply clarify
- defer
- apply clarify
- apply (subgoal_tac "k = j")
- apply (simp_all (no_asm_simp))
- apply (case_tac "Suc (k + l) = j")
- apply (subgoal_tac "funsum f k l = 0")
- apply (rule_tac [2] funsum_zero)
- apply (subgoal_tac [3] "f (Suc (k + l)) = 0")
- apply (subgoal_tac [3] "j \<le> k + l")
- prefer 4
- apply arith
- apply auto
- done
-
-
-subsection {* Chinese: uniqueness *}
-
-lemma zcong_funprod_aux:
- "m_cond n mf ==> km_cond n kf mf
- ==> lincong_sol n kf bf mf x ==> lincong_sol n kf bf mf y
- ==> [x = y] (mod mf n)"
- apply (unfold m_cond_def km_cond_def lincong_sol_def)
- apply (rule iffD1)
- apply (rule_tac k = "kf n" in zcong_cancel2)
- apply (rule_tac [3] b = "bf n" in zcong_trans)
- prefer 4
- apply (subst zcong_sym)
- defer
- apply (rule order_less_imp_le)
- apply simp_all
- done
-
-lemma zcong_funprod [rule_format]:
- "m_cond n mf --> km_cond n kf mf -->
- lincong_sol n kf bf mf x --> lincong_sol n kf bf mf y -->
- [x = y] (mod funprod mf 0 n)"
- apply (induct n)
- apply (simp_all (no_asm))
- apply (blast intro: zcong_funprod_aux)
- apply (rule impI)+
- apply (rule zcong_zgcd_zmult_zmod)
- apply (blast intro: zcong_funprod_aux)
- prefer 2
- apply (subst zgcd_commute)
- apply (rule funprod_zgcd)
- apply (auto simp add: m_cond_def km_cond_def lincong_sol_def)
- done
-
-
-subsection {* Chinese: existence *}
-
-lemma unique_xi_sol:
- "0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf
- ==> \<exists>!x. 0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
- apply (rule zcong_lineq_unique)
- apply (tactic {* stac (thm "zgcd_zmult_cancel") 2 *})
- apply (unfold m_cond_def km_cond_def mhf_def)
- apply (simp_all (no_asm_simp))
- apply safe
- apply (tactic {* stac (thm "zgcd_zmult_cancel") 3 *})
- apply (rule_tac [!] funprod_zgcd)
- apply safe
- apply simp_all
- apply (subgoal_tac "i<n")
- prefer 2
- apply arith
- apply (case_tac [2] i)
- apply simp_all
- done
-
-lemma x_sol_lin_aux:
- "0 < n ==> i \<le> n ==> j \<le> n ==> j \<noteq> i ==> mf j dvd mhf mf n i"
- apply (unfold mhf_def)
- apply (case_tac "i = 0")
- apply (case_tac [2] "i = n")
- apply (simp_all (no_asm_simp))
- apply (case_tac [3] "j < i")
- apply (rule_tac [3] dvd_mult2)
- apply (rule_tac [4] dvd_mult)
- apply (rule_tac [!] funprod_zdvd)
- apply arith
- apply arith
- apply arith
- apply arith
- apply arith
- apply arith
- apply arith
- apply arith
- done
-
-lemma x_sol_lin:
- "0 < n ==> i \<le> n
- ==> x_sol n kf bf mf mod mf i =
- xilin_sol i n kf bf mf * mhf mf n i mod mf i"
- apply (unfold x_sol_def)
- apply (subst funsum_mod)
- apply (subst funsum_oneelem)
- apply auto
- apply (subst dvd_eq_mod_eq_0 [symmetric])
- apply (rule dvd_mult)
- apply (rule x_sol_lin_aux)
- apply auto
- done
-
-
-subsection {* Chinese *}
-
-lemma chinese_remainder:
- "0 < n ==> m_cond n mf ==> km_cond n kf mf
- ==> \<exists>!x. 0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x"
- apply safe
- apply (rule_tac [2] m = "funprod mf 0 n" in zcong_zless_imp_eq)
- apply (rule_tac [6] zcong_funprod)
- apply auto
- apply (rule_tac x = "x_sol n kf bf mf mod funprod mf 0 n" in exI)
- apply (unfold lincong_sol_def)
- apply safe
- apply (tactic {* stac (thm "zcong_zmod") 3 *})
- apply (tactic {* stac (thm "mod_mult_eq") 3 *})
- apply (tactic {* stac (thm "mod_mod_cancel") 3 *})
- apply (tactic {* stac (thm "x_sol_lin") 4 *})
- apply (tactic {* stac (thm "mod_mult_eq" RS sym) 6 *})
- apply (tactic {* stac (thm "zcong_zmod" RS sym) 6 *})
- apply (subgoal_tac [6]
- "0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
- \<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)")
- prefer 6
- apply (simp add: zmult_ac)
- apply (unfold xilin_sol_def)
- apply (tactic {* asm_simp_tac @{simpset} 6 *})
- apply (rule_tac [6] ex1_implies_ex [THEN someI_ex])
- apply (rule_tac [6] unique_xi_sol)
- apply (rule_tac [3] funprod_zdvd)
- apply (unfold m_cond_def)
- apply (rule funprod_pos [THEN pos_mod_sign])
- apply (rule_tac [2] funprod_pos [THEN pos_mod_bound])
- apply auto
- done
-
-end
--- a/src/HOL/NumberTheory/Euler.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,304 +0,0 @@
-(* Title: HOL/Quadratic_Reciprocity/Euler.thy
- ID: $Id$
- Authors: Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {* Euler's criterion *}
-
-theory Euler imports Residues EvenOdd begin
-
-definition
- MultInvPair :: "int => int => int => int set" where
- "MultInvPair a p j = {StandardRes p j, StandardRes p (a * (MultInv p j))}"
-
-definition
- SetS :: "int => int => int set set" where
- "SetS a p = (MultInvPair a p ` SRStar p)"
-
-
-subsection {* Property for MultInvPair *}
-
-lemma MultInvPair_prop1a:
- "[| zprime p; 2 < p; ~([a = 0](mod p));
- X \<in> (SetS a p); Y \<in> (SetS a p);
- ~((X \<inter> Y) = {}) |] ==> X = Y"
- apply (auto simp add: SetS_def)
- apply (drule StandardRes_SRStar_prop1a)+ defer 1
- apply (drule StandardRes_SRStar_prop1a)+
- apply (auto simp add: MultInvPair_def StandardRes_prop2 zcong_sym)
- apply (drule notE, rule MultInv_zcong_prop1, auto)[]
- apply (drule notE, rule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
- apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
- apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
- apply (drule MultInv_zcong_prop1, auto)[]
- apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
- apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
- apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
- done
-
-lemma MultInvPair_prop1b:
- "[| zprime p; 2 < p; ~([a = 0](mod p));
- X \<in> (SetS a p); Y \<in> (SetS a p);
- X \<noteq> Y |] ==> X \<inter> Y = {}"
- apply (rule notnotD)
- apply (rule notI)
- apply (drule MultInvPair_prop1a, auto)
- done
-
-lemma MultInvPair_prop1c: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==>
- \<forall>X \<in> SetS a p. \<forall>Y \<in> SetS a p. X \<noteq> Y --> X\<inter>Y = {}"
- by (auto simp add: MultInvPair_prop1b)
-
-lemma MultInvPair_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==>
- Union ( SetS a p) = SRStar p"
- apply (auto simp add: SetS_def MultInvPair_def StandardRes_SRStar_prop4
- SRStar_mult_prop2)
- apply (frule StandardRes_SRStar_prop3)
- apply (rule bexI, auto)
- done
-
-lemma MultInvPair_distinct: "[| zprime p; 2 < p; ~([a = 0] (mod p));
- ~([j = 0] (mod p));
- ~(QuadRes p a) |] ==>
- ~([j = a * MultInv p j] (mod p))"
-proof
- assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and
- "~([j = 0] (mod p))" and "~(QuadRes p a)"
- assume "[j = a * MultInv p j] (mod p)"
- then have "[j * j = (a * MultInv p j) * j] (mod p)"
- by (auto simp add: zcong_scalar)
- then have a:"[j * j = a * (MultInv p j * j)] (mod p)"
- by (auto simp add: zmult_ac)
- have "[j * j = a] (mod p)"
- proof -
- from prems have b: "[MultInv p j * j = 1] (mod p)"
- by (simp add: MultInv_prop2a)
- from b a show ?thesis
- by (auto simp add: zcong_zmult_prop2)
- qed
- then have "[j^2 = a] (mod p)"
- by (metis number_of_is_id power2_eq_square succ_bin_simps)
- with prems show False
- by (simp add: QuadRes_def)
-qed
-
-lemma MultInvPair_card_two: "[| zprime p; 2 < p; ~([a = 0] (mod p));
- ~(QuadRes p a); ~([j = 0] (mod p)) |] ==>
- card (MultInvPair a p j) = 2"
- apply (auto simp add: MultInvPair_def)
- apply (subgoal_tac "~ (StandardRes p j = StandardRes p (a * MultInv p j))")
- apply auto
- apply (metis MultInvPair_distinct Pls_def StandardRes_def aux number_of_is_id one_is_num_one)
- done
-
-
-subsection {* Properties of SetS *}
-
-lemma SetS_finite: "2 < p ==> finite (SetS a p)"
- by (auto simp add: SetS_def SRStar_finite [of p] finite_imageI)
-
-lemma SetS_elems_finite: "\<forall>X \<in> SetS a p. finite X"
- by (auto simp add: SetS_def MultInvPair_def)
-
-lemma SetS_elems_card: "[| zprime p; 2 < p; ~([a = 0] (mod p));
- ~(QuadRes p a) |] ==>
- \<forall>X \<in> SetS a p. card X = 2"
- apply (auto simp add: SetS_def)
- apply (frule StandardRes_SRStar_prop1a)
- apply (rule MultInvPair_card_two, auto)
- done
-
-lemma Union_SetS_finite: "2 < p ==> finite (Union (SetS a p))"
- by (auto simp add: SetS_finite SetS_elems_finite finite_Union)
-
-lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set);
- \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S"
- by (induct set: finite) auto
-
-lemma SetS_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
- int(card(SetS a p)) = (p - 1) div 2"
-proof -
- assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)"
- then have "(p - 1) = 2 * int(card(SetS a p))"
- proof -
- have "p - 1 = int(card(Union (SetS a p)))"
- by (auto simp add: prems MultInvPair_prop2 SRStar_card)
- also have "... = int (setsum card (SetS a p))"
- by (auto simp add: prems SetS_finite SetS_elems_finite
- MultInvPair_prop1c [of p a] card_Union_disjoint)
- also have "... = int(setsum (%x.2) (SetS a p))"
- using prems
- by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite
- card_setsum_aux simp del: setsum_constant)
- also have "... = 2 * int(card( SetS a p))"
- by (auto simp add: prems SetS_finite setsum_const2)
- finally show ?thesis .
- qed
- from this show ?thesis
- by auto
-qed
-
-lemma SetS_setprod_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p));
- ~(QuadRes p a); x \<in> (SetS a p) |] ==>
- [\<Prod>x = a] (mod p)"
- apply (auto simp add: SetS_def MultInvPair_def)
- apply (frule StandardRes_SRStar_prop1a)
- apply (subgoal_tac "StandardRes p x \<noteq> StandardRes p (a * MultInv p x)")
- apply (auto simp add: StandardRes_prop2 MultInvPair_distinct)
- apply (frule_tac m = p and x = x and y = "(a * MultInv p x)" in
- StandardRes_prop4)
- apply (subgoal_tac "[x * (a * MultInv p x) = a * (x * MultInv p x)] (mod p)")
- apply (drule_tac a = "StandardRes p x * StandardRes p (a * MultInv p x)" and
- b = "x * (a * MultInv p x)" and
- c = "a * (x * MultInv p x)" in zcong_trans, force)
- apply (frule_tac p = p and x = x in MultInv_prop2, auto)
-apply (metis StandardRes_SRStar_prop3 mult_1_right mult_commute zcong_sym zcong_zmult_prop1)
- apply (auto simp add: zmult_ac)
- done
-
-lemma aux1: "[| 0 < x; (x::int) < a; x \<noteq> (a - 1) |] ==> x < a - 1"
- by arith
-
-lemma aux2: "[| (a::int) < c; b < c |] ==> (a \<le> b | b \<le> a)"
- by auto
-
-lemma SRStar_d22set_prop: "2 < p \<Longrightarrow> (SRStar p) = {1} \<union> (d22set (p - 1))"
- apply (induct p rule: d22set.induct)
- apply auto
- apply (simp add: SRStar_def d22set.simps)
- apply (simp add: SRStar_def d22set.simps, clarify)
- apply (frule aux1)
- apply (frule aux2, auto)
- apply (simp_all add: SRStar_def)
- apply (simp add: d22set.simps)
- apply (frule d22set_le)
- apply (frule d22set_g_1, auto)
- done
-
-lemma Union_SetS_setprod_prop1: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
- [\<Prod>(Union (SetS a p)) = a ^ nat ((p - 1) div 2)] (mod p)"
-proof -
- assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and "~(QuadRes p a)"
- then have "[\<Prod>(Union (SetS a p)) =
- setprod (setprod (%x. x)) (SetS a p)] (mod p)"
- by (auto simp add: SetS_finite SetS_elems_finite
- MultInvPair_prop1c setprod_Union_disjoint)
- also have "[setprod (setprod (%x. x)) (SetS a p) =
- setprod (%x. a) (SetS a p)] (mod p)"
- by (rule setprod_same_function_zcong)
- (auto simp add: prems SetS_setprod_prop SetS_finite)
- also (zcong_trans) have "[setprod (%x. a) (SetS a p) =
- a^(card (SetS a p))] (mod p)"
- by (auto simp add: prems SetS_finite setprod_constant)
- finally (zcong_trans) show ?thesis
- apply (rule zcong_trans)
- apply (subgoal_tac "card(SetS a p) = nat((p - 1) div 2)", auto)
- apply (subgoal_tac "nat(int(card(SetS a p))) = nat((p - 1) div 2)", force)
- apply (auto simp add: prems SetS_card)
- done
-qed
-
-lemma Union_SetS_setprod_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==>
- \<Prod>(Union (SetS a p)) = zfact (p - 1)"
-proof -
- assume "zprime p" and "2 < p" and "~([a = 0](mod p))"
- then have "\<Prod>(Union (SetS a p)) = \<Prod>(SRStar p)"
- by (auto simp add: MultInvPair_prop2)
- also have "... = \<Prod>({1} \<union> (d22set (p - 1)))"
- by (auto simp add: prems SRStar_d22set_prop)
- also have "... = zfact(p - 1)"
- proof -
- have "~(1 \<in> d22set (p - 1)) & finite( d22set (p - 1))"
- by (metis d22set_fin d22set_g_1 linorder_neq_iff)
- then have "\<Prod>({1} \<union> (d22set (p - 1))) = \<Prod>(d22set (p - 1))"
- by auto
- then show ?thesis
- by (auto simp add: d22set_prod_zfact)
- qed
- finally show ?thesis .
-qed
-
-lemma zfact_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
- [zfact (p - 1) = a ^ nat ((p - 1) div 2)] (mod p)"
- apply (frule Union_SetS_setprod_prop1)
- apply (auto simp add: Union_SetS_setprod_prop2)
- done
-
-text {* \medskip Prove the first part of Euler's Criterion: *}
-
-lemma Euler_part1: "[| 2 < p; zprime p; ~([x = 0](mod p));
- ~(QuadRes p x) |] ==>
- [x^(nat (((p) - 1) div 2)) = -1](mod p)"
- by (metis Wilson_Russ number_of_is_id zcong_sym zcong_trans zfact_prop)
-
-text {* \medskip Prove another part of Euler Criterion: *}
-
-lemma aux_1: "0 < p ==> (a::int) ^ nat (p) = a * a ^ (nat (p) - 1)"
-proof -
- assume "0 < p"
- then have "a ^ (nat p) = a ^ (1 + (nat p - 1))"
- by (auto simp add: diff_add_assoc)
- also have "... = (a ^ 1) * a ^ (nat(p) - 1)"
- by (simp only: zpower_zadd_distrib)
- also have "... = a * a ^ (nat(p) - 1)"
- by auto
- finally show ?thesis .
-qed
-
-lemma aux_2: "[| (2::int) < p; p \<in> zOdd |] ==> 0 < ((p - 1) div 2)"
-proof -
- assume "2 < p" and "p \<in> zOdd"
- then have "(p - 1):zEven"
- by (auto simp add: zEven_def zOdd_def)
- then have aux_1: "2 * ((p - 1) div 2) = (p - 1)"
- by (auto simp add: even_div_2_prop2)
- with `2 < p` have "1 < (p - 1)"
- by auto
- then have " 1 < (2 * ((p - 1) div 2))"
- by (auto simp add: aux_1)
- then have "0 < (2 * ((p - 1) div 2)) div 2"
- by auto
- then show ?thesis by auto
-qed
-
-lemma Euler_part2:
- "[| 2 < p; zprime p; [a = 0] (mod p) |] ==> [0 = a ^ nat ((p - 1) div 2)] (mod p)"
- apply (frule zprime_zOdd_eq_grt_2)
- apply (frule aux_2, auto)
- apply (frule_tac a = a in aux_1, auto)
- apply (frule zcong_zmult_prop1, auto)
- done
-
-text {* \medskip Prove the final part of Euler's Criterion: *}
-
-lemma aux__1: "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> ~(p dvd y)"
- by (metis dvdI power2_eq_square zcong_sym zcong_trans zcong_zero_equiv_div dvd_trans)
-
-lemma aux__2: "2 * nat((p - 1) div 2) = nat (2 * ((p - 1) div 2))"
- by (auto simp add: nat_mult_distrib)
-
-lemma Euler_part3: "[| 2 < p; zprime p; ~([x = 0](mod p)); QuadRes p x |] ==>
- [x^(nat (((p) - 1) div 2)) = 1](mod p)"
- apply (subgoal_tac "p \<in> zOdd")
- apply (auto simp add: QuadRes_def)
- prefer 2
- apply (metis number_of_is_id numeral_1_eq_1 zprime_zOdd_eq_grt_2)
- apply (frule aux__1, auto)
- apply (drule_tac z = "nat ((p - 1) div 2)" in zcong_zpower)
- apply (auto simp add: zpower_zpower)
- apply (rule zcong_trans)
- apply (auto simp add: zcong_sym [of "x ^ nat ((p - 1) div 2)"])
- apply (metis Little_Fermat even_div_2_prop2 mult_Bit0 number_of_is_id odd_minus_one_even one_is_num_one zmult_1 aux__2)
- done
-
-
-text {* \medskip Finally show Euler's Criterion: *}
-
-theorem Euler_Criterion: "[| 2 < p; zprime p |] ==> [(Legendre a p) =
- a^(nat (((p) - 1) div 2))] (mod p)"
- apply (auto simp add: Legendre_def Euler_part2)
- apply (frule Euler_part3, auto simp add: zcong_sym)[]
- apply (frule Euler_part1, auto simp add: zcong_sym)[]
- done
-
-end
--- a/src/HOL/NumberTheory/EulerFermat.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,348 +0,0 @@
-(* Title: HOL/NumberTheory/EulerFermat.thy
- ID: $Id$
- Author: Thomas M. Rasmussen
- Copyright 2000 University of Cambridge
-*)
-
-header {* Fermat's Little Theorem extended to Euler's Totient function *}
-
-theory EulerFermat
-imports BijectionRel IntFact
-begin
-
-text {*
- Fermat's Little Theorem extended to Euler's Totient function. More
- abstract approach than Boyer-Moore (which seems necessary to achieve
- the extended version).
-*}
-
-
-subsection {* Definitions and lemmas *}
-
-inductive_set
- RsetR :: "int => int set set"
- for m :: int
- where
- empty [simp]: "{} \<in> RsetR m"
- | insert: "A \<in> RsetR m ==> zgcd a m = 1 ==>
- \<forall>a'. a' \<in> A --> \<not> zcong a a' m ==> insert a A \<in> RsetR m"
-
-consts
- BnorRset :: "int * int => int set"
-
-recdef BnorRset
- "measure ((\<lambda>(a, m). nat a) :: int * int => nat)"
- "BnorRset (a, m) =
- (if 0 < a then
- let na = BnorRset (a - 1, m)
- in (if zgcd a m = 1 then insert a na else na)
- else {})"
-
-definition
- norRRset :: "int => int set" where
- "norRRset m = BnorRset (m - 1, m)"
-
-definition
- noXRRset :: "int => int => int set" where
- "noXRRset m x = (\<lambda>a. a * x) ` norRRset m"
-
-definition
- phi :: "int => nat" where
- "phi m = card (norRRset m)"
-
-definition
- is_RRset :: "int set => int => bool" where
- "is_RRset A m = (A \<in> RsetR m \<and> card A = phi m)"
-
-definition
- RRset2norRR :: "int set => int => int => int" where
- "RRset2norRR A m a =
- (if 1 < m \<and> is_RRset A m \<and> a \<in> A then
- SOME b. zcong a b m \<and> b \<in> norRRset m
- else 0)"
-
-definition
- zcongm :: "int => int => int => bool" where
- "zcongm m = (\<lambda>a b. zcong a b m)"
-
-lemma abs_eq_1_iff [iff]: "(abs z = (1::int)) = (z = 1 \<or> z = -1)"
- -- {* LCP: not sure why this lemma is needed now *}
- by (auto simp add: abs_if)
-
-
-text {* \medskip @{text norRRset} *}
-
-declare BnorRset.simps [simp del]
-
-lemma BnorRset_induct:
- assumes "!!a m. P {} a m"
- and "!!a m. 0 < (a::int) ==> P (BnorRset (a - 1, m::int)) (a - 1) m
- ==> P (BnorRset(a,m)) a m"
- shows "P (BnorRset(u,v)) u v"
- apply (rule BnorRset.induct)
- apply safe
- apply (case_tac [2] "0 < a")
- apply (rule_tac [2] prems)
- apply simp_all
- apply (simp_all add: BnorRset.simps prems)
- done
-
-lemma Bnor_mem_zle [rule_format]: "b \<in> BnorRset (a, m) \<longrightarrow> b \<le> a"
- apply (induct a m rule: BnorRset_induct)
- apply simp
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- done
-
-lemma Bnor_mem_zle_swap: "a < b ==> b \<notin> BnorRset (a, m)"
- by (auto dest: Bnor_mem_zle)
-
-lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> 0 < b"
- apply (induct a m rule: BnorRset_induct)
- prefer 2
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- done
-
-lemma Bnor_mem_if [rule_format]:
- "zgcd b m = 1 --> 0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
- apply (induct a m rule: BnorRset.induct, auto)
- apply (subst BnorRset.simps)
- defer
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- done
-
-lemma Bnor_in_RsetR [rule_format]: "a < m --> BnorRset (a, m) \<in> RsetR m"
- apply (induct a m rule: BnorRset_induct, simp)
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- apply (rule RsetR.insert)
- apply (rule_tac [3] allI)
- apply (rule_tac [3] impI)
- apply (rule_tac [3] zcong_not)
- apply (subgoal_tac [6] "a' \<le> a - 1")
- apply (rule_tac [7] Bnor_mem_zle)
- apply (rule_tac [5] Bnor_mem_zg, auto)
- done
-
-lemma Bnor_fin: "finite (BnorRset (a, m))"
- apply (induct a m rule: BnorRset_induct)
- prefer 2
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- done
-
-lemma norR_mem_unique_aux: "a \<le> b - 1 ==> a < (b::int)"
- apply auto
- done
-
-lemma norR_mem_unique:
- "1 < m ==>
- zgcd a m = 1 ==> \<exists>!b. [a = b] (mod m) \<and> b \<in> norRRset m"
- apply (unfold norRRset_def)
- apply (cut_tac a = a and m = m in zcong_zless_unique, auto)
- apply (rule_tac [2] m = m in zcong_zless_imp_eq)
- apply (auto intro: Bnor_mem_zle Bnor_mem_zg zcong_trans
- order_less_imp_le norR_mem_unique_aux simp add: zcong_sym)
- apply (rule_tac x = b in exI, safe)
- apply (rule Bnor_mem_if)
- apply (case_tac [2] "b = 0")
- apply (auto intro: order_less_le [THEN iffD2])
- prefer 2
- apply (simp only: zcong_def)
- apply (subgoal_tac "zgcd a m = m")
- prefer 2
- apply (subst zdvd_iff_zgcd [symmetric])
- apply (rule_tac [4] zgcd_zcong_zgcd)
- apply (simp_all add: zcong_sym)
- done
-
-
-text {* \medskip @{term noXRRset} *}
-
-lemma RRset_gcd [rule_format]:
- "is_RRset A m ==> a \<in> A --> zgcd a m = 1"
- apply (unfold is_RRset_def)
- apply (rule RsetR.induct [where P="%A. a \<in> A --> zgcd a m = 1"], auto)
- done
-
-lemma RsetR_zmult_mono:
- "A \<in> RsetR m ==>
- 0 < m ==> zgcd x m = 1 ==> (\<lambda>a. a * x) ` A \<in> RsetR m"
- apply (erule RsetR.induct, simp_all)
- apply (rule RsetR.insert, auto)
- apply (blast intro: zgcd_zgcd_zmult)
- apply (simp add: zcong_cancel)
- done
-
-lemma card_nor_eq_noX:
- "0 < m ==>
- zgcd x m = 1 ==> card (noXRRset m x) = card (norRRset m)"
- apply (unfold norRRset_def noXRRset_def)
- apply (rule card_image)
- apply (auto simp add: inj_on_def Bnor_fin)
- apply (simp add: BnorRset.simps)
- done
-
-lemma noX_is_RRset:
- "0 < m ==> zgcd x m = 1 ==> is_RRset (noXRRset m x) m"
- apply (unfold is_RRset_def phi_def)
- apply (auto simp add: card_nor_eq_noX)
- apply (unfold noXRRset_def norRRset_def)
- apply (rule RsetR_zmult_mono)
- apply (rule Bnor_in_RsetR, simp_all)
- done
-
-lemma aux_some:
- "1 < m ==> is_RRset A m ==> a \<in> A
- ==> zcong a (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) m \<and>
- (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) \<in> norRRset m"
- apply (rule norR_mem_unique [THEN ex1_implies_ex, THEN someI_ex])
- apply (rule_tac [2] RRset_gcd, simp_all)
- done
-
-lemma RRset2norRR_correct:
- "1 < m ==> is_RRset A m ==> a \<in> A ==>
- [a = RRset2norRR A m a] (mod m) \<and> RRset2norRR A m a \<in> norRRset m"
- apply (unfold RRset2norRR_def, simp)
- apply (rule aux_some, simp_all)
- done
-
-lemmas RRset2norRR_correct1 =
- RRset2norRR_correct [THEN conjunct1, standard]
-lemmas RRset2norRR_correct2 =
- RRset2norRR_correct [THEN conjunct2, standard]
-
-lemma RsetR_fin: "A \<in> RsetR m ==> finite A"
- by (induct set: RsetR) auto
-
-lemma RRset_zcong_eq [rule_format]:
- "1 < m ==>
- is_RRset A m ==> [a = b] (mod m) ==> a \<in> A --> b \<in> A --> a = b"
- apply (unfold is_RRset_def)
- apply (rule RsetR.induct [where P="%A. a \<in> A --> b \<in> A --> a = b"])
- apply (auto simp add: zcong_sym)
- done
-
-lemma aux:
- "P (SOME a. P a) ==> Q (SOME a. Q a) ==>
- (SOME a. P a) = (SOME a. Q a) ==> \<exists>a. P a \<and> Q a"
- apply auto
- done
-
-lemma RRset2norRR_inj:
- "1 < m ==> is_RRset A m ==> inj_on (RRset2norRR A m) A"
- apply (unfold RRset2norRR_def inj_on_def, auto)
- apply (subgoal_tac "\<exists>b. ([x = b] (mod m) \<and> b \<in> norRRset m) \<and>
- ([y = b] (mod m) \<and> b \<in> norRRset m)")
- apply (rule_tac [2] aux)
- apply (rule_tac [3] aux_some)
- apply (rule_tac [2] aux_some)
- apply (rule RRset_zcong_eq, auto)
- apply (rule_tac b = b in zcong_trans)
- apply (simp_all add: zcong_sym)
- done
-
-lemma RRset2norRR_eq_norR:
- "1 < m ==> is_RRset A m ==> RRset2norRR A m ` A = norRRset m"
- apply (rule card_seteq)
- prefer 3
- apply (subst card_image)
- apply (rule_tac RRset2norRR_inj, auto)
- apply (rule_tac [3] RRset2norRR_correct2, auto)
- apply (unfold is_RRset_def phi_def norRRset_def)
- apply (auto simp add: Bnor_fin)
- done
-
-
-lemma Bnor_prod_power_aux: "a \<notin> A ==> inj f ==> f a \<notin> f ` A"
-by (unfold inj_on_def, auto)
-
-lemma Bnor_prod_power [rule_format]:
- "x \<noteq> 0 ==> a < m --> \<Prod>((\<lambda>a. a * x) ` BnorRset (a, m)) =
- \<Prod>(BnorRset(a, m)) * x^card (BnorRset (a, m))"
- apply (induct a m rule: BnorRset_induct)
- prefer 2
- apply (simplesubst BnorRset.simps) --{*multiple redexes*}
- apply (unfold Let_def, auto)
- apply (simp add: Bnor_fin Bnor_mem_zle_swap)
- apply (subst setprod_insert)
- apply (rule_tac [2] Bnor_prod_power_aux)
- apply (unfold inj_on_def)
- apply (simp_all add: zmult_ac Bnor_fin finite_imageI
- Bnor_mem_zle_swap)
- done
-
-
-subsection {* Fermat *}
-
-lemma bijzcong_zcong_prod:
- "(A, B) \<in> bijR (zcongm m) ==> [\<Prod>A = \<Prod>B] (mod m)"
- apply (unfold zcongm_def)
- apply (erule bijR.induct)
- apply (subgoal_tac [2] "a \<notin> A \<and> b \<notin> B \<and> finite A \<and> finite B")
- apply (auto intro: fin_bijRl fin_bijRr zcong_zmult)
- done
-
-lemma Bnor_prod_zgcd [rule_format]:
- "a < m --> zgcd (\<Prod>(BnorRset(a, m))) m = 1"
- apply (induct a m rule: BnorRset_induct)
- prefer 2
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto)
- apply (simp add: Bnor_fin Bnor_mem_zle_swap)
- apply (blast intro: zgcd_zgcd_zmult)
- done
-
-theorem Euler_Fermat:
- "0 < m ==> zgcd x m = 1 ==> [x^(phi m) = 1] (mod m)"
- apply (unfold norRRset_def phi_def)
- apply (case_tac "x = 0")
- apply (case_tac [2] "m = 1")
- apply (rule_tac [3] iffD1)
- apply (rule_tac [3] k = "\<Prod>(BnorRset(m - 1, m))"
- in zcong_cancel2)
- prefer 5
- apply (subst Bnor_prod_power [symmetric])
- apply (rule_tac [7] Bnor_prod_zgcd, simp_all)
- apply (rule bijzcong_zcong_prod)
- apply (fold norRRset_def noXRRset_def)
- apply (subst RRset2norRR_eq_norR [symmetric])
- apply (rule_tac [3] inj_func_bijR, auto)
- apply (unfold zcongm_def)
- apply (rule_tac [2] RRset2norRR_correct1)
- apply (rule_tac [5] RRset2norRR_inj)
- apply (auto intro: order_less_le [THEN iffD2]
- simp add: noX_is_RRset)
- apply (unfold noXRRset_def norRRset_def)
- apply (rule finite_imageI)
- apply (rule Bnor_fin)
- done
-
-lemma Bnor_prime:
- "\<lbrakk> zprime p; a < p \<rbrakk> \<Longrightarrow> card (BnorRset (a, p)) = nat a"
- apply (induct a p rule: BnorRset.induct)
- apply (subst BnorRset.simps)
- apply (unfold Let_def, auto simp add:zless_zprime_imp_zrelprime)
- apply (subgoal_tac "finite (BnorRset (a - 1,m))")
- apply (subgoal_tac "a ~: BnorRset (a - 1,m)")
- apply (auto simp add: card_insert_disjoint Suc_nat_eq_nat_zadd1)
- apply (frule Bnor_mem_zle, arith)
- apply (frule Bnor_fin)
- done
-
-lemma phi_prime: "zprime p ==> phi p = nat (p - 1)"
- apply (unfold phi_def norRRset_def)
- apply (rule Bnor_prime, auto)
- done
-
-theorem Little_Fermat:
- "zprime p ==> \<not> p dvd x ==> [x^(nat (p - 1)) = 1] (mod p)"
- apply (subst phi_prime [symmetric])
- apply (rule_tac [2] Euler_Fermat)
- apply (erule_tac [3] zprime_imp_zrelprime)
- apply (unfold zprime_def, auto)
- done
-
-end
--- a/src/HOL/NumberTheory/EvenOdd.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,256 +0,0 @@
-(* Title: HOL/Quadratic_Reciprocity/EvenOdd.thy
- Authors: Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {*Parity: Even and Odd Integers*}
-
-theory EvenOdd
-imports Int2
-begin
-
-definition
- zOdd :: "int set" where
- "zOdd = {x. \<exists>k. x = 2 * k + 1}"
-
-definition
- zEven :: "int set" where
- "zEven = {x. \<exists>k. x = 2 * k}"
-
-subsection {* Some useful properties about even and odd *}
-
-lemma zOddI [intro?]: "x = 2 * k + 1 \<Longrightarrow> x \<in> zOdd"
- and zOddE [elim?]: "x \<in> zOdd \<Longrightarrow> (!!k. x = 2 * k + 1 \<Longrightarrow> C) \<Longrightarrow> C"
- by (auto simp add: zOdd_def)
-
-lemma zEvenI [intro?]: "x = 2 * k \<Longrightarrow> x \<in> zEven"
- and zEvenE [elim?]: "x \<in> zEven \<Longrightarrow> (!!k. x = 2 * k \<Longrightarrow> C) \<Longrightarrow> C"
- by (auto simp add: zEven_def)
-
-lemma one_not_even: "~(1 \<in> zEven)"
-proof
- assume "1 \<in> zEven"
- then obtain k :: int where "1 = 2 * k" ..
- then show False by arith
-qed
-
-lemma even_odd_conj: "~(x \<in> zOdd & x \<in> zEven)"
-proof -
- {
- fix a b
- assume "2 * (a::int) = 2 * (b::int) + 1"
- then have "2 * (a::int) - 2 * (b :: int) = 1"
- by arith
- then have "2 * (a - b) = 1"
- by (auto simp add: zdiff_zmult_distrib)
- moreover have "(2 * (a - b)):zEven"
- by (auto simp only: zEven_def)
- ultimately have False
- by (auto simp add: one_not_even)
- }
- then show ?thesis
- by (auto simp add: zOdd_def zEven_def)
-qed
-
-lemma even_odd_disj: "(x \<in> zOdd | x \<in> zEven)"
- by (simp add: zOdd_def zEven_def) arith
-
-lemma not_odd_impl_even: "~(x \<in> zOdd) ==> x \<in> zEven"
- using even_odd_disj by auto
-
-lemma odd_mult_odd_prop: "(x*y):zOdd ==> x \<in> zOdd"
-proof (rule classical)
- assume "\<not> ?thesis"
- then have "x \<in> zEven" by (rule not_odd_impl_even)
- then obtain a where a: "x = 2 * a" ..
- assume "x * y : zOdd"
- then obtain b where "x * y = 2 * b + 1" ..
- with a have "2 * a * y = 2 * b + 1" by simp
- then have "2 * a * y - 2 * b = 1"
- by arith
- then have "2 * (a * y - b) = 1"
- by (auto simp add: zdiff_zmult_distrib)
- moreover have "(2 * (a * y - b)):zEven"
- by (auto simp only: zEven_def)
- ultimately have False
- by (auto simp add: one_not_even)
- then show ?thesis ..
-qed
-
-lemma odd_minus_one_even: "x \<in> zOdd ==> (x - 1):zEven"
- by (auto simp add: zOdd_def zEven_def)
-
-lemma even_div_2_prop1: "x \<in> zEven ==> (x mod 2) = 0"
- by (auto simp add: zEven_def)
-
-lemma even_div_2_prop2: "x \<in> zEven ==> (2 * (x div 2)) = x"
- by (auto simp add: zEven_def)
-
-lemma even_plus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x + y \<in> zEven"
- apply (auto simp add: zEven_def)
- apply (auto simp only: zadd_zmult_distrib2 [symmetric])
- done
-
-lemma even_times_either: "x \<in> zEven ==> x * y \<in> zEven"
- by (auto simp add: zEven_def)
-
-lemma even_minus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x - y \<in> zEven"
- apply (auto simp add: zEven_def)
- apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
- done
-
-lemma odd_minus_odd: "[| x \<in> zOdd; y \<in> zOdd |] ==> x - y \<in> zEven"
- apply (auto simp add: zOdd_def zEven_def)
- apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
- done
-
-lemma even_minus_odd: "[| x \<in> zEven; y \<in> zOdd |] ==> x - y \<in> zOdd"
- apply (auto simp add: zOdd_def zEven_def)
- apply (rule_tac x = "k - ka - 1" in exI)
- apply auto
- done
-
-lemma odd_minus_even: "[| x \<in> zOdd; y \<in> zEven |] ==> x - y \<in> zOdd"
- apply (auto simp add: zOdd_def zEven_def)
- apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
- done
-
-lemma odd_times_odd: "[| x \<in> zOdd; y \<in> zOdd |] ==> x * y \<in> zOdd"
- apply (auto simp add: zOdd_def zadd_zmult_distrib zadd_zmult_distrib2)
- apply (rule_tac x = "2 * ka * k + ka + k" in exI)
- apply (auto simp add: zadd_zmult_distrib)
- done
-
-lemma odd_iff_not_even: "(x \<in> zOdd) = (~ (x \<in> zEven))"
- using even_odd_conj even_odd_disj by auto
-
-lemma even_product: "x * y \<in> zEven ==> x \<in> zEven | y \<in> zEven"
- using odd_iff_not_even odd_times_odd by auto
-
-lemma even_diff: "x - y \<in> zEven = ((x \<in> zEven) = (y \<in> zEven))"
-proof
- assume xy: "x - y \<in> zEven"
- {
- assume x: "x \<in> zEven"
- have "y \<in> zEven"
- proof (rule classical)
- assume "\<not> ?thesis"
- then have "y \<in> zOdd"
- by (simp add: odd_iff_not_even)
- with x have "x - y \<in> zOdd"
- by (simp add: even_minus_odd)
- with xy have False
- by (auto simp add: odd_iff_not_even)
- then show ?thesis ..
- qed
- } moreover {
- assume y: "y \<in> zEven"
- have "x \<in> zEven"
- proof (rule classical)
- assume "\<not> ?thesis"
- then have "x \<in> zOdd"
- by (auto simp add: odd_iff_not_even)
- with y have "x - y \<in> zOdd"
- by (simp add: odd_minus_even)
- with xy have False
- by (auto simp add: odd_iff_not_even)
- then show ?thesis ..
- qed
- }
- ultimately show "(x \<in> zEven) = (y \<in> zEven)"
- by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
- even_minus_odd odd_minus_even)
-next
- assume "(x \<in> zEven) = (y \<in> zEven)"
- then show "x - y \<in> zEven"
- by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
- even_minus_odd odd_minus_even)
-qed
-
-lemma neg_one_even_power: "[| x \<in> zEven; 0 \<le> x |] ==> (-1::int)^(nat x) = 1"
-proof -
- assume "x \<in> zEven" and "0 \<le> x"
- from `x \<in> zEven` obtain a where "x = 2 * a" ..
- with `0 \<le> x` have "0 \<le> a" by simp
- from `0 \<le> x` and `x = 2 * a` have "nat x = nat (2 * a)"
- by simp
- also from `x = 2 * a` have "nat (2 * a) = 2 * nat a"
- by (simp add: nat_mult_distrib)
- finally have "(-1::int)^nat x = (-1)^(2 * nat a)"
- by simp
- also have "... = ((-1::int)^2)^ (nat a)"
- by (simp add: zpower_zpower [symmetric])
- also have "(-1::int)^2 = 1"
- by simp
- finally show ?thesis
- by simp
-qed
-
-lemma neg_one_odd_power: "[| x \<in> zOdd; 0 \<le> x |] ==> (-1::int)^(nat x) = -1"
-proof -
- assume "x \<in> zOdd" and "0 \<le> x"
- from `x \<in> zOdd` obtain a where "x = 2 * a + 1" ..
- with `0 \<le> x` have a: "0 \<le> a" by simp
- with `0 \<le> x` and `x = 2 * a + 1` have "nat x = nat (2 * a + 1)"
- by simp
- also from a have "nat (2 * a + 1) = 2 * nat a + 1"
- by (auto simp add: nat_mult_distrib nat_add_distrib)
- finally have "(-1::int)^nat x = (-1)^(2 * nat a + 1)"
- by simp
- also have "... = ((-1::int)^2)^ (nat a) * (-1)^1"
- by (auto simp add: zpower_zpower [symmetric] zpower_zadd_distrib)
- also have "(-1::int)^2 = 1"
- by simp
- finally show ?thesis
- by simp
-qed
-
-lemma neg_one_power_parity: "[| 0 \<le> x; 0 \<le> y; (x \<in> zEven) = (y \<in> zEven) |] ==>
- (-1::int)^(nat x) = (-1::int)^(nat y)"
- using even_odd_disj [of x] even_odd_disj [of y]
- by (auto simp add: neg_one_even_power neg_one_odd_power)
-
-
-lemma one_not_neg_one_mod_m: "2 < m ==> ~([1 = -1] (mod m))"
- by (auto simp add: zcong_def zdvd_not_zless)
-
-lemma even_div_2_l: "[| y \<in> zEven; x < y |] ==> x div 2 < y div 2"
-proof -
- assume "y \<in> zEven" and "x < y"
- from `y \<in> zEven` obtain k where k: "y = 2 * k" ..
- with `x < y` have "x < 2 * k" by simp
- then have "x div 2 < k" by (auto simp add: div_prop1)
- also have "k = (2 * k) div 2" by simp
- finally have "x div 2 < 2 * k div 2" by simp
- with k show ?thesis by simp
-qed
-
-lemma even_sum_div_2: "[| x \<in> zEven; y \<in> zEven |] ==> (x + y) div 2 = x div 2 + y div 2"
- by (auto simp add: zEven_def)
-
-lemma even_prod_div_2: "[| x \<in> zEven |] ==> (x * y) div 2 = (x div 2) * y"
- by (auto simp add: zEven_def)
-
-(* An odd prime is greater than 2 *)
-
-lemma zprime_zOdd_eq_grt_2: "zprime p ==> (p \<in> zOdd) = (2 < p)"
- apply (auto simp add: zOdd_def zprime_def)
- apply (drule_tac x = 2 in allE)
- using odd_iff_not_even [of p]
- apply (auto simp add: zOdd_def zEven_def)
- done
-
-(* Powers of -1 and parity *)
-
-lemma neg_one_special: "finite A ==>
- ((-1 :: int) ^ card A) * (-1 ^ card A) = 1"
- by (induct set: finite) auto
-
-lemma neg_one_power: "(-1::int)^n = 1 | (-1::int)^n = -1"
- by (induct n) auto
-
-lemma neg_one_power_eq_mod_m: "[| 2 < m; [(-1::int)^j = (-1::int)^k] (mod m) |]
- ==> ((-1::int)^j = (-1::int)^k)"
- using neg_one_power [of j] and ListMem.insert neg_one_power [of k]
- by (auto simp add: one_not_neg_one_mod_m zcong_sym)
-
-end
--- a/src/HOL/NumberTheory/Factorization.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,341 +0,0 @@
-(* Title: HOL/NumberTheory/Factorization.thy
- ID: $Id$
- Author: Thomas Marthedal Rasmussen
- Copyright 2000 University of Cambridge
-*)
-
-header {* Fundamental Theorem of Arithmetic (unique factorization into primes) *}
-
-theory Factorization
-imports Main Primes Permutation
-begin
-
-
-subsection {* Definitions *}
-
-definition
- primel :: "nat list => bool" where
- "primel xs = (\<forall>p \<in> set xs. prime p)"
-
-consts
- nondec :: "nat list => bool "
- prod :: "nat list => nat"
- oinsert :: "nat => nat list => nat list"
- sort :: "nat list => nat list"
-
-primrec
- "nondec [] = True"
- "nondec (x # xs) = (case xs of [] => True | y # ys => x \<le> y \<and> nondec xs)"
-
-primrec
- "prod [] = Suc 0"
- "prod (x # xs) = x * prod xs"
-
-primrec
- "oinsert x [] = [x]"
- "oinsert x (y # ys) = (if x \<le> y then x # y # ys else y # oinsert x ys)"
-
-primrec
- "sort [] = []"
- "sort (x # xs) = oinsert x (sort xs)"
-
-
-subsection {* Arithmetic *}
-
-lemma one_less_m: "(m::nat) \<noteq> m * k ==> m \<noteq> Suc 0 ==> Suc 0 < m"
- apply (cases m)
- apply auto
- done
-
-lemma one_less_k: "(m::nat) \<noteq> m * k ==> Suc 0 < m * k ==> Suc 0 < k"
- apply (cases k)
- apply auto
- done
-
-lemma mult_left_cancel: "(0::nat) < k ==> k * n = k * m ==> n = m"
- apply auto
- done
-
-lemma mn_eq_m_one: "(0::nat) < m ==> m * n = m ==> n = Suc 0"
- apply (cases n)
- apply auto
- done
-
-lemma prod_mn_less_k:
- "(0::nat) < n ==> 0 < k ==> Suc 0 < m ==> m * n = k ==> n < k"
- apply (induct m)
- apply auto
- done
-
-
-subsection {* Prime list and product *}
-
-lemma prod_append: "prod (xs @ ys) = prod xs * prod ys"
- apply (induct xs)
- apply (simp_all add: mult_assoc)
- done
-
-lemma prod_xy_prod:
- "prod (x # xs) = prod (y # ys) ==> x * prod xs = y * prod ys"
- apply auto
- done
-
-lemma primel_append: "primel (xs @ ys) = (primel xs \<and> primel ys)"
- apply (unfold primel_def)
- apply auto
- done
-
-lemma prime_primel: "prime n ==> primel [n] \<and> prod [n] = n"
- apply (unfold primel_def)
- apply auto
- done
-
-lemma prime_nd_one: "prime p ==> \<not> p dvd Suc 0"
- apply (unfold prime_def dvd_def)
- apply auto
- done
-
-lemma hd_dvd_prod: "prod (x # xs) = prod ys ==> x dvd (prod ys)"
- by (metis dvd_mult_left dvd_refl prod.simps(2))
-
-lemma primel_tl: "primel (x # xs) ==> primel xs"
- apply (unfold primel_def)
- apply auto
- done
-
-lemma primel_hd_tl: "(primel (x # xs)) = (prime x \<and> primel xs)"
- apply (unfold primel_def)
- apply auto
- done
-
-lemma primes_eq: "prime p ==> prime q ==> p dvd q ==> p = q"
- apply (unfold prime_def)
- apply auto
- done
-
-lemma primel_one_empty: "primel xs ==> prod xs = Suc 0 ==> xs = []"
- apply (cases xs)
- apply (simp_all add: primel_def prime_def)
- done
-
-lemma prime_g_one: "prime p ==> Suc 0 < p"
- apply (unfold prime_def)
- apply auto
- done
-
-lemma prime_g_zero: "prime p ==> 0 < p"
- apply (unfold prime_def)
- apply auto
- done
-
-lemma primel_nempty_g_one:
- "primel xs \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> Suc 0 < prod xs"
- apply (induct xs)
- apply simp
- apply (fastsimp simp: primel_def prime_def elim: one_less_mult)
- done
-
-lemma primel_prod_gz: "primel xs ==> 0 < prod xs"
- apply (induct xs)
- apply (auto simp: primel_def prime_def)
- done
-
-
-subsection {* Sorting *}
-
-lemma nondec_oinsert: "nondec xs \<Longrightarrow> nondec (oinsert x xs)"
- apply (induct xs)
- apply simp
- apply (case_tac xs)
- apply (simp_all cong del: list.weak_case_cong)
- done
-
-lemma nondec_sort: "nondec (sort xs)"
- apply (induct xs)
- apply simp_all
- apply (erule nondec_oinsert)
- done
-
-lemma x_less_y_oinsert: "x \<le> y ==> l = y # ys ==> x # l = oinsert x l"
- apply simp_all
- done
-
-lemma nondec_sort_eq [rule_format]: "nondec xs \<longrightarrow> xs = sort xs"
- apply (induct xs)
- apply safe
- apply simp_all
- apply (case_tac xs)
- apply simp_all
- apply (case_tac xs)
- apply simp
- apply (rule_tac y = aa and ys = list in x_less_y_oinsert)
- apply simp_all
- done
-
-lemma oinsert_x_y: "oinsert x (oinsert y l) = oinsert y (oinsert x l)"
- apply (induct l)
- apply auto
- done
-
-
-subsection {* Permutation *}
-
-lemma perm_primel [rule_format]: "xs <~~> ys ==> primel xs --> primel ys"
- apply (unfold primel_def)
- apply (induct set: perm)
- apply simp
- apply simp
- apply (simp (no_asm))
- apply blast
- apply blast
- done
-
-lemma perm_prod: "xs <~~> ys ==> prod xs = prod ys"
- apply (induct set: perm)
- apply (simp_all add: mult_ac)
- done
-
-lemma perm_subst_oinsert: "xs <~~> ys ==> oinsert a xs <~~> oinsert a ys"
- apply (induct set: perm)
- apply auto
- done
-
-lemma perm_oinsert: "x # xs <~~> oinsert x xs"
- apply (induct xs)
- apply auto
- done
-
-lemma perm_sort: "xs <~~> sort xs"
- apply (induct xs)
- apply (auto intro: perm_oinsert elim: perm_subst_oinsert)
- done
-
-lemma perm_sort_eq: "xs <~~> ys ==> sort xs = sort ys"
- apply (induct set: perm)
- apply (simp_all add: oinsert_x_y)
- done
-
-
-subsection {* Existence *}
-
-lemma ex_nondec_lemma:
- "primel xs ==> \<exists>ys. primel ys \<and> nondec ys \<and> prod ys = prod xs"
- apply (blast intro: nondec_sort perm_prod perm_primel perm_sort perm_sym)
- done
-
-lemma not_prime_ex_mk:
- "Suc 0 < n \<and> \<not> prime n ==>
- \<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k"
- apply (unfold prime_def dvd_def)
- apply (auto intro: n_less_m_mult_n n_less_n_mult_m one_less_m one_less_k)
- done
-
-lemma split_primel:
- "primel xs \<Longrightarrow> primel ys \<Longrightarrow> \<exists>l. primel l \<and> prod l = prod xs * prod ys"
- apply (rule exI)
- apply safe
- apply (rule_tac [2] prod_append)
- apply (simp add: primel_append)
- done
-
-lemma factor_exists [rule_format]: "Suc 0 < n --> (\<exists>l. primel l \<and> prod l = n)"
- apply (induct n rule: nat_less_induct)
- apply (rule impI)
- apply (case_tac "prime n")
- apply (rule exI)
- apply (erule prime_primel)
- apply (cut_tac n = n in not_prime_ex_mk)
- apply (auto intro!: split_primel)
- done
-
-lemma nondec_factor_exists: "Suc 0 < n ==> \<exists>l. primel l \<and> nondec l \<and> prod l = n"
- apply (erule factor_exists [THEN exE])
- apply (blast intro!: ex_nondec_lemma)
- done
-
-
-subsection {* Uniqueness *}
-
-lemma prime_dvd_mult_list [rule_format]:
- "prime p ==> p dvd (prod xs) --> (\<exists>m. m:set xs \<and> p dvd m)"
- apply (induct xs)
- apply (force simp add: prime_def)
- apply (force dest: prime_dvd_mult)
- done
-
-lemma hd_xs_dvd_prod:
- "primel (x # xs) ==> primel ys ==> prod (x # xs) = prod ys
- ==> \<exists>m. m \<in> set ys \<and> x dvd m"
- apply (rule prime_dvd_mult_list)
- apply (simp add: primel_hd_tl)
- apply (erule hd_dvd_prod)
- done
-
-lemma prime_dvd_eq: "primel (x # xs) ==> primel ys ==> m \<in> set ys ==> x dvd m ==> x = m"
- apply (rule primes_eq)
- apply (auto simp add: primel_def primel_hd_tl)
- done
-
-lemma hd_xs_eq_prod:
- "primel (x # xs) ==>
- primel ys ==> prod (x # xs) = prod ys ==> x \<in> set ys"
- apply (frule hd_xs_dvd_prod)
- apply auto
- apply (drule prime_dvd_eq)
- apply auto
- done
-
-lemma perm_primel_ex:
- "primel (x # xs) ==>
- primel ys ==> prod (x # xs) = prod ys ==> \<exists>l. ys <~~> (x # l)"
- apply (rule exI)
- apply (rule perm_remove)
- apply (erule hd_xs_eq_prod)
- apply simp_all
- done
-
-lemma primel_prod_less:
- "primel (x # xs) ==>
- primel ys ==> prod (x # xs) = prod ys ==> prod xs < prod ys"
- by (metis less_asym linorder_neqE_nat mult_less_cancel2 nat_0_less_mult_iff
- nat_less_le nat_mult_1 prime_def primel_hd_tl primel_prod_gz prod.simps(2))
-
-lemma prod_one_empty:
- "primel xs ==> p * prod xs = p ==> prime p ==> xs = []"
- apply (auto intro: primel_one_empty simp add: prime_def)
- done
-
-lemma uniq_ex_aux:
- "\<forall>m. m < prod ys --> (\<forall>xs ys. primel xs \<and> primel ys \<and>
- prod xs = prod ys \<and> prod xs = m --> xs <~~> ys) ==>
- primel list ==> primel x ==> prod list = prod x ==> prod x < prod ys
- ==> x <~~> list"
- apply simp
- done
-
-lemma factor_unique [rule_format]:
- "\<forall>xs ys. primel xs \<and> primel ys \<and> prod xs = prod ys \<and> prod xs = n
- --> xs <~~> ys"
- apply (induct n rule: nat_less_induct)
- apply safe
- apply (case_tac xs)
- apply (force intro: primel_one_empty)
- apply (rule perm_primel_ex [THEN exE])
- apply simp_all
- apply (rule perm.trans [THEN perm_sym])
- apply assumption
- apply (rule perm.Cons)
- apply (case_tac "x = []")
- apply (metis perm_prod perm_refl prime_primel primel_hd_tl primel_tl prod_one_empty)
- apply (metis nat_0_less_mult_iff nat_mult_eq_cancel1 perm_primel perm_prod primel_prod_gz primel_prod_less primel_tl prod.simps(2))
- done
-
-lemma perm_nondec_unique:
- "xs <~~> ys ==> nondec xs ==> nondec ys ==> xs = ys"
- by (metis nondec_sort_eq perm_sort_eq)
-
-theorem unique_prime_factorization [rule_format]:
- "\<forall>n. Suc 0 < n --> (\<exists>!l. primel l \<and> nondec l \<and> prod l = n)"
- by (metis factor_unique nondec_factor_exists perm_nondec_unique)
-
-end
--- a/src/HOL/NumberTheory/Fib.thy Tue Sep 01 14:13:34 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,150 +0,0 @@
-(* ID: $Id$
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
- Copyright 1997 University of Cambridge
-*)
-
-header {* The Fibonacci function *}
-
-theory Fib
-imports Primes
-begin
-
-text {*
- Fibonacci numbers: proofs of laws taken from:
- R. L. Graham, D. E. Knuth, O. Patashnik. Concrete Mathematics.
- (Addison-Wesley, 1989)
-
- \bigskip
-*}
-
-fun fib :: "nat \<Rightarrow> nat"
-where
- "fib 0 = 0"
-| "fib (Suc 0) = 1"
-| fib_2: "fib (Suc (Suc n)) = fib n + fib (Suc n)"
-
-text {*
- \medskip The difficulty in these proofs is to ensure that the
- induction hypotheses are applied before the definition of @{term
- fib}. Towards this end, the @{term fib} equations are not declared
- to the Simplifier and are applied very selectively at first.
-*}
-
-text{*We disable @{text fib.fib_2fib_2} for simplification ...*}
-declare fib_2 [simp del]
-
-text{*...then prove a version that has a more restrictive pattern.*}
-lemma fib_Suc3: "fib (Suc (Suc (Suc n))) = fib (Suc n) + fib (Suc (Suc n))"
- by (rule fib_2)
-
-text {* \medskip Concrete Mathematics, page 280 *}
-
-lemma fib_add: "fib (Suc (n + k)) = fib (Suc k) * fib (Suc n) + fib k * fib n"
-proof (induct n rule: fib.induct)
- case 1 show ?case by simp
-next
- case 2 show ?case by (simp add: fib_2)
-next
- case 3 thus ?case by (simp add: fib_2 add_mult_distrib2)
-qed
-
-lemma fib_Suc_neq_0: "fib (Suc n) \<noteq> 0"
- apply (induct n rule: fib.induct)
- apply (simp_all add: fib_2)
- done
-
-lemma fib_Suc_gr_0: "0 < fib (Suc n)"
- by (insert fib_Suc_neq_0 [of n], simp)
-
-lemma fib_gr_0: "0 < n ==> 0 < fib n"
- by (case_tac n, auto simp add: fib_Suc_gr_0)
-
-
-text {*
- \medskip Concrete Mathematics, page 278: Cassini's identity. The proof is
- much easier using integers, not natural numbers!
-*}
-
-lemma fib_Cassini_int:
- "int (fib (Suc (Suc n)) * fib n) =
- (if n mod 2 = 0 then int (fib (Suc n) * fib (Suc n)) - 1
- else int (fib (Suc n) * fib (Suc n)) + 1)"
-proof(induct n rule: fib.induct)
- case 1 thus ?case by (simp add: fib_2)
-next
- case 2 thus ?case by (simp add: fib_2 mod_Suc)
-next
- case (3 x)
- have "Suc 0 \<noteq> x mod 2 \<longrightarrow> x mod 2 = 0" by presburger
- with "3.hyps" show ?case by (simp add: fib.simps add_mult_distrib add_mult_distrib2)
-qed
-
-text{*We now obtain a version for the natural numbers via the coercion
- function @{term int}.*}
-theorem fib_Cassini:
- "fib (Suc (Suc n)) * fib n =
- (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1
- else fib (Suc n) * fib (Suc n) + 1)"
- apply (rule int_int_eq [THEN iffD1])
- apply (simp add: fib_Cassini_int)
- apply (subst zdiff_int [symmetric])
- apply (insert fib_Suc_gr_0 [of n], simp_all)
- done
-
-
-text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
-
-lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (Suc n)) = Suc 0"
- apply (induct n rule: fib.induct)
- prefer 3
- apply (simp add: gcd_commute fib_Suc3)
- apply (simp_all add: fib_2)
- done
-
-lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)"
- apply (simp add: gcd_commute [of "fib m"])
- apply (case_tac m)
- apply simp
- apply (simp add: fib_add)
- apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0])
- apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric])
- apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel)
- done
-
-lemma gcd_fib_diff: "m \<le> n ==> gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
- by (simp add: gcd_fib_add [symmetric, of _ "n-m"])
-
-lemma gcd_fib_mod: "0 < m ==> gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-proof (induct n rule: less_induct)
- case (less n)
- from less.prems have pos_m: "0 < m" .
- show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
- proof (cases "m < n")
- case True note m_n = True
- then have m_n': "m \<le> n" by auto
- with pos_m have pos_n: "0 < n" by auto
- with pos_m m_n have diff: "n - m < n" by auto
- have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
- by (simp add: mod_if [of n]) (insert m_n, auto)
- also have "\<dots> = gcd (fib m) (fib (n - m))" by (simp add: less.hyps diff pos_m)
- also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff m_n')