src/HOL/Computational_Algebra/Primes.thy
author haftmann
Sun, 02 Nov 2025 20:01:43 +0100
changeset 83359 518a1464f1ac
parent 83358 1cf4f1e930af
permissions -rw-r--r--
more efficient naive prime test
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65435
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     1
(*  Title:      HOL/Computational_Algebra/Primes.thy
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     2
    Author:     Christophe Tabacznyj
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     3
    Author:     Lawrence C. Paulson
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     4
    Author:     Amine Chaieb
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     5
    Author:     Thomas M. Rasmussen
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     6
    Author:     Jeremy Avigad
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     7
    Author:     Tobias Nipkow
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     8
    Author:     Manuel Eberl
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
     9
65435
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
    10
This theory deals with properties of primes. Definitions and lemmas are
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
    11
proved uniformly for the natural numbers and integers.
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    12
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    13
This file combines and revises a number of prior developments.
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    14
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    15
The original theories "GCD" and "Primes" were by Christophe Tabacznyj
58623
2db1df2c8467 more bibtex entries;
wenzelm
parents: 57512
diff changeset
    16
and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    17
gcd, lcm, and prime for the natural numbers.
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    18
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    19
The original theory "IntPrimes" was by Thomas M. Rasmussen, and
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    20
extended gcd, lcm, primes to the integers. Amine Chaieb provided
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    21
another extension of the notions to the integers, and added a number
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    22
of results to "Primes" and "GCD". IntPrimes also defined and developed
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    23
the congruence relations on the integers. The notion was extended to
33718
06e9aff51d17 Fixed a typo in a comment.
webertj
parents: 32479
diff changeset
    24
the natural numbers by Chaieb.
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    25
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
    26
Jeremy Avigad combined all of these, made everything uniform for the
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
    27
natural numbers and the integers, and added a number of new theorems.
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
    28
31798
fe9a3043d36c Cleaned up GCD
nipkow
parents: 31766
diff changeset
    29
Tobias Nipkow cleaned up a lot.
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    30
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    31
Florian Haftmann and Manuel Eberl put primality and prime factorisation
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    32
onto an algebraic foundation and thus generalised these concepts to 
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    33
other rings, such as polynomials. (see also the Factorial_Ring theory).
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    34
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    35
There were also previous formalisations of unique factorisation by 
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    36
Thomas Marthedal Rasmussen, Jeremy Avigad, and David Gray.
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    37
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    38
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
    39
section \<open>Primes\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    40
32479
521cc9bf2958 some reorganization of number theory
haftmann
parents: 32415
diff changeset
    41
theory Primes
68623
b942da0962c2 correct import
nipkow
parents: 67613
diff changeset
    42
imports Euclidean_Algorithm
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    43
begin
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
    44
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69313
diff changeset
    45
subsection \<open>Primes on \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close>\<close>
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    46
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    47
lemma Suc_0_not_prime_nat [simp]: "\<not> prime (Suc 0)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    48
  using not_prime_1 [where ?'a = nat] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    49
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    50
lemma prime_ge_2_nat:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    51
  "p \<ge> 2" if "prime p" for p :: nat
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    52
proof -
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    53
  from that have "p \<noteq> 0" and "p \<noteq> 1"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    54
    by (auto dest: prime_elem_not_zeroI prime_elem_not_unit)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    55
  then show ?thesis
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    56
    by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    57
qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    58
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    59
lemma prime_ge_2_int:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    60
  "p \<ge> 2" if "prime p" for p :: int
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    61
proof -
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    62
  from that have "prime_elem p" and "\<bar>p\<bar> = p"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    63
    by (auto dest: normalize_prime)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    64
  then have "p \<noteq> 0" and "\<bar>p\<bar> \<noteq> 1" and "p \<ge> 0"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    65
    by (auto dest: prime_elem_not_zeroI prime_elem_not_unit)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    66
  then show ?thesis
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    67
    by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    68
qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    69
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    70
lemma prime_ge_0_int: "prime p \<Longrightarrow> p \<ge> (0::int)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    71
  using prime_ge_2_int [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    72
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    73
lemma prime_gt_0_nat: "prime p \<Longrightarrow> p > (0::nat)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    74
  using prime_ge_2_nat [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    75
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    76
(* As a simp or intro rule,
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    77
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    78
     prime p \<Longrightarrow> p > 0
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    79
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    80
   wreaks havoc here. When the premise includes \<forall>x \<in># M. prime x, it
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    81
   leads to the backchaining
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    82
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    83
     x > 0
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    84
     prime x
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    85
     x \<in># M   which is, unfortunately,
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    86
     count M x > 0  FIXME no, this is obsolete
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    87
*)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
    88
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    89
lemma prime_gt_0_int: "prime p \<Longrightarrow> p > (0::int)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    90
  using prime_ge_2_int [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    91
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    92
lemma prime_ge_1_nat: "prime p \<Longrightarrow> p \<ge> (1::nat)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    93
  using prime_ge_2_nat [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    94
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    95
lemma prime_ge_Suc_0_nat: "prime p \<Longrightarrow> p \<ge> Suc 0"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    96
  using prime_ge_1_nat [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    97
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    98
lemma prime_ge_1_int: "prime p \<Longrightarrow> p \<ge> (1::int)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
    99
  using prime_ge_2_int [of p] by simp
31706
1db0c8f235fb new GCD library, courtesy of Jeremy Avigad
huffman
parents: 30738
diff changeset
   100
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   101
lemma prime_gt_1_nat: "prime p \<Longrightarrow> p > (1::nat)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   102
  using prime_ge_2_nat [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   103
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   104
lemma prime_gt_Suc_0_nat: "prime p \<Longrightarrow> p > Suc 0"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   105
  using prime_gt_1_nat [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   106
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   107
lemma prime_gt_1_int: "prime p \<Longrightarrow> p > (1::int)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   108
  using prime_ge_2_int [of p] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   109
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   110
lemma prime_natI:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   111
  "prime p" if "p \<ge> 2" and "\<And>m n. p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n" for p :: nat
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   112
  using that by (auto intro!: primeI prime_elemI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   113
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   114
lemma prime_intI:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   115
  "prime p" if "p \<ge> 2" and "\<And>m n. p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n" for p :: int
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   116
  using that by (auto intro!: primeI prime_elemI)
66837
6ba663ff2b1c tuned proofs
haftmann
parents: 66453
diff changeset
   117
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   118
lemma prime_elem_nat_iff [simp]:
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   119
  "prime_elem n \<longleftrightarrow> prime n" for n :: nat
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   120
  by (simp add: prime_def)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   121
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   122
lemma prime_elem_iff_prime_abs [simp]:
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   123
  "prime_elem k \<longleftrightarrow> prime \<bar>k\<bar>" for k :: int
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   124
  by (auto intro: primeI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   125
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   126
lemma prime_nat_int_transfer [simp]:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   127
  "prime (int n) \<longleftrightarrow> prime n" (is "?P \<longleftrightarrow> ?Q")
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   128
proof
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   129
  assume ?P
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   130
  then have "n \<ge> 2"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   131
    by (auto dest: prime_ge_2_int)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   132
  then show ?Q
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   133
  proof (rule prime_natI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   134
    fix r s
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   135
    assume "n dvd r * s"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   136
    with of_nat_dvd_iff [of n "r * s"] have "int n dvd int r * int s"
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   137
      by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   138
    with \<open>?P\<close> have "int n dvd int r \<or> int n dvd int s"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   139
      using prime_dvd_mult_iff [of "int n" "int r" "int s"]
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   140
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   141
    then show "n dvd r \<or> n dvd s"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   142
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   143
  qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   144
next
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   145
  assume ?Q
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   146
  then have "int n \<ge> 2"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   147
    by (auto dest: prime_ge_2_nat)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   148
  then show ?P
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   149
  proof (rule prime_intI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   150
    fix r s
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   151
    assume "int n dvd r * s"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   152
    then have "n dvd nat \<bar>r * s\<bar>"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   153
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   154
    then have "n dvd nat \<bar>r\<bar> * nat \<bar>s\<bar>"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   155
      by (simp add: nat_abs_mult_distrib)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   156
    with \<open>?Q\<close> have "n dvd nat \<bar>r\<bar> \<or> n dvd nat \<bar>s\<bar>"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   157
      using prime_dvd_mult_iff [of "n" "nat \<bar>r\<bar>" "nat \<bar>s\<bar>"]
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   158
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   159
    then show "int n dvd r \<or> int n dvd s"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   160
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   161
  qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   162
qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   163
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   164
lemma prime_nat_iff_prime [simp]:
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   165
  "prime (nat k) \<longleftrightarrow> prime k"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   166
proof (cases "k \<ge> 0")
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   167
  case True
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   168
  then show ?thesis
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   169
    using prime_nat_int_transfer [of "nat k"] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   170
next
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   171
  case False
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   172
  then show ?thesis
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   173
    by (auto dest: prime_ge_2_int)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   174
qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   175
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   176
lemma prime_int_nat_transfer:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   177
  "prime k \<longleftrightarrow> k \<ge> 0 \<and> prime (nat k)"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   178
  by (auto dest: prime_ge_2_int)
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   179
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   180
lemma prime_nat_naiveI:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   181
  "prime p" if "p \<ge> 2" and dvd: "\<And>n. n dvd p \<Longrightarrow> n = 1 \<or> n = p" for p :: nat
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   182
proof (rule primeI, rule prime_elemI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   183
  fix m n :: nat
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   184
  assume "p dvd m * n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   185
  then obtain r s where "p = r * s" "r dvd m" "s dvd n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   186
    by (blast dest: division_decomp)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   187
  moreover have "r = 1 \<or> r = p"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   188
    using \<open>r dvd m\<close> \<open>p = r * s\<close> dvd [of r] by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   189
  ultimately show "p dvd m \<or> p dvd n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   190
    by auto
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   191
qed (use \<open>p \<ge> 2\<close> in simp_all)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   192
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   193
lemma prime_int_naiveI:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   194
  "prime p" if "p \<ge> 2" and dvd: "\<And>k. k dvd p \<Longrightarrow> \<bar>k\<bar> = 1 \<or> \<bar>k\<bar> = p" for p :: int
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   195
proof -
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   196
  from \<open>p \<ge> 2\<close> have "nat p \<ge> 2"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   197
    by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   198
  then have "prime (nat p)"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   199
  proof (rule prime_nat_naiveI)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   200
    fix n
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   201
    assume "n dvd nat p"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   202
    with \<open>p \<ge> 2\<close> have "n dvd nat \<bar>p\<bar>"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   203
      by simp
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   204
    then have "int n dvd p"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   205
      by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   206
    with dvd [of "int n"] show "n = 1 \<or> n = nat p"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   207
      by auto
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   208
  qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   209
  then show ?thesis
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   210
    by simp
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   211
qed
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   212
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   213
lemma prime_nat_iff:
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   214
  "prime (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   215
proof (safe intro!: prime_gt_1_nat)
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   216
  assume "prime n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   217
  then have *: "prime_elem n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   218
    by simp
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   219
  fix m assume m: "m dvd n" "m \<noteq> n"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   220
  from * \<open>m dvd n\<close> have "n dvd m \<or> is_unit m"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   221
    by (intro irreducibleD' prime_elem_imp_irreducible)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   222
  with m show "m = 1" by (auto dest: dvd_antisym)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   223
next
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   224
  assume "n > 1" "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n"
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   225
  then show "prime n"
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   226
    using prime_nat_naiveI [of n] by auto
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   227
qed
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   228
83359
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   229
lemma prime_nat_iff':
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   230
  "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. \<not> n dvd p)"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   231
proof safe
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   232
  assume "p > 1" and *: "\<forall>n\<in>{2..<p}. \<not>n dvd p"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   233
  show "prime p" unfolding prime_nat_iff
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   234
  proof (intro conjI allI impI)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   235
    fix m assume "m dvd p"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   236
    with \<open>p > 1\<close> have "m \<noteq> 0" by (intro notI) auto
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   237
    hence "m \<ge> 1" by simp
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   238
    moreover from \<open>m dvd p\<close> and * have "m \<notin> {2..<p}" by blast
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   239
    with \<open>m dvd p\<close> and \<open>p > 1\<close> have "m \<le> 1 \<or> m = p" by (auto dest: dvd_imp_le)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   240
    ultimately show "m = 1 \<or> m = p" by simp
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   241
  qed fact+
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   242
qed (auto simp: prime_nat_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   243
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   244
lemma prime_int_iff:
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   245
  "prime (n::int) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n))"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   246
proof (intro iffI conjI allI impI; (elim conjE)?)
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   247
  assume *: "prime n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   248
  hence irred: "irreducible n" by (auto intro: prime_elem_imp_irreducible)
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   249
  from * have "n \<ge> 0" "n \<noteq> 0" "n \<noteq> 1"
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   250
    by (auto simp add: prime_ge_0_int)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   251
  thus "n > 1" by presburger
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   252
  fix m assume "m dvd n" \<open>m \<ge> 0\<close>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   253
  with irred have "m dvd 1 \<or> n dvd m" by (auto simp: irreducible_altdef)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   254
  with \<open>m dvd n\<close> \<open>m \<ge> 0\<close> \<open>n > 1\<close> show "m = 1 \<or> m = n"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   255
    using associated_iff_dvd[of m n] by auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   256
next
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   257
  assume n: "1 < n" "\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   258
  hence "nat n > 1" by simp
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   259
  moreover have "\<forall>m. m dvd nat n \<longrightarrow> m = 1 \<or> m = nat n"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   260
  proof (intro allI impI)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   261
    fix m assume "m dvd nat n"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   262
    with \<open>n > 1\<close> have "m dvd nat \<bar>n\<bar>"
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   263
      by simp
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   264
    then have "int m dvd n"
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   265
      by simp
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65435
diff changeset
   266
    with n(2) have "int m = 1 \<or> int m = n"
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65435
diff changeset
   267
      using of_nat_0_le_iff by blast
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   268
    thus "m = 1 \<or> m = nat n" by auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   269
  qed
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   270
  ultimately show "prime n" 
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   271
    unfolding prime_int_nat_transfer prime_nat_iff by auto
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   272
qed
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   273
83359
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   274
lemma prime_int_iff':
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   275
  "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. \<not> n dvd p)" (is "?P \<longleftrightarrow> ?Q")
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   276
proof (cases "p \<ge> 0")
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   277
  case True
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   278
  have "?P \<longleftrightarrow> prime (nat p)"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   279
    by simp
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   280
  also have "\<dots> \<longleftrightarrow> p > 1 \<and> (\<forall>n\<in>{2..<nat p}. \<not> n dvd nat \<bar>p\<bar>)"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   281
    using True by (simp add: prime_nat_iff')
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   282
  also have "{2..<nat p} = nat ` {2..<p}"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   283
    using True int_eq_iff by fastforce
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   284
  finally show "?P \<longleftrightarrow> ?Q" by simp
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   285
next
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   286
  case False
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   287
  then show ?thesis
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   288
    by (auto simp add: prime_ge_0_int)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   289
qed
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   290
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   291
lemma prime_nat_not_dvd:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   292
  assumes "prime p" "p > n" "n \<noteq> (1::nat)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   293
  shows   "\<not>n dvd p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   294
proof
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   295
  assume "n dvd p"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   296
  from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   297
  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   298
    by (cases "n = 0") (auto dest!: dvd_imp_le)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   299
qed
22027
e4a08629c4bd A few lemmas about relative primes when dividing trough gcd
chaieb
parents: 21404
diff changeset
   300
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   301
lemma prime_int_not_dvd:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   302
  assumes "prime p" "p > n" "n > (1::int)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   303
  shows   "\<not>n dvd p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   304
proof
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   305
  assume "n dvd p"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   306
  from assms(1) have "irreducible p" by (auto intro: prime_elem_imp_irreducible)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   307
  from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   308
    by (auto dest!: zdvd_imp_le)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   309
qed
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   310
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   311
lemma prime_odd_nat: "prime p \<Longrightarrow> p > (2::nat) \<Longrightarrow> odd p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   312
  by (intro prime_nat_not_dvd) auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   313
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   314
lemma prime_odd_int: "prime p \<Longrightarrow> p > (2::int) \<Longrightarrow> odd p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   315
  by (intro prime_int_not_dvd) auto
22367
6860f09242bf tuned document;
wenzelm
parents: 22027
diff changeset
   316
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   317
lemma prime_int_altdef:
55242
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55238
diff changeset
   318
  "prime p = (1 < p \<and> (\<forall>m::int. m \<ge> 0 \<longrightarrow> m dvd p \<longrightarrow>
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55238
diff changeset
   319
    m = 1 \<or> m = p))"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   320
  unfolding prime_int_iff by blast
27568
9949dc7a24de Theorem names as in IntPrimes.thy, also several theorems moved from there
chaieb
parents: 27556
diff changeset
   321
62481
b5d8e57826df tuned bootstrap order to provide type classes in a more sensible order
haftmann
parents: 62429
diff changeset
   322
lemma not_prime_eq_prod_nat:
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   323
  assumes "m > 1" "\<not> prime (m::nat)"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   324
  shows   "\<exists>n k. n = m * k \<and> 1 < m \<and> m < n \<and> 1 < k \<and> k < n"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   325
  using assms irreducible_altdef[of m]
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   326
  by (auto simp: prime_elem_iff_irreducible irreducible_altdef)
53598
2bd8d459ebae remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents: 47108
diff changeset
   327
83359
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   328
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   329
subsection \<open>Make prime naively executable\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   330
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   331
lemma prime_int_numeral_eq [simp]:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   332
  "prime (numeral m :: int) \<longleftrightarrow> prime (numeral m :: nat)"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   333
  by (simp add: prime_int_nat_transfer)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   334
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   335
class check_prime_by_range = normalization_semidom + discrete_linordered_semidom +
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   336
  assumes prime_iff: \<open>prime a \<longleftrightarrow> 1 < a \<and> (\<forall>d\<in>{2..a div 2}. \<not> d dvd a)\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   337
begin
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   338
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   339
lemma two_is_prime [simp]:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   340
  \<open>prime 2\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   341
  by (simp add: prime_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   342
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   343
end
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   344
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   345
lemma divisor_less_eq_half_nat:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   346
  \<open>m \<le> n div 2\<close> if \<open>m dvd n\<close> \<open>m < n\<close> for m n :: nat
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   347
  using that by (auto simp add: less_eq_div_iff_mult_less_eq)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   348
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   349
instance nat :: check_prime_by_range
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   350
  apply standard
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   351
  apply (auto simp add: prime_nat_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   352
  apply (rule ccontr)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   353
  apply (auto simp add: neq_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   354
  apply (metis One_nat_def Suc_1 Suc_leI atLeastAtMost_iff divisor_less_eq_half_nat)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   355
  done
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   356
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   357
lemma two_is_prime_nat [simp]:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   358
  \<open>prime (2::nat)\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   359
  by (fact two_is_prime)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   360
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   361
lemma divisor_less_eq_half_int:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   362
  \<open>k \<le> l div 2\<close> if \<open>k dvd l\<close> \<open>k < l\<close> \<open>l \<ge> 0\<close> \<open>k \<ge> 0\<close> for k l :: int
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   363
proof -
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   364
  define m n where \<open>m = nat \<bar>k\<bar>\<close> \<open>n = nat \<bar>l\<bar>\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   365
  with \<open>l \<ge> 0\<close> \<open>k \<ge> 0\<close> have \<open>k = int m\<close> \<open>l = int n\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   366
    by simp_all
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   367
  with that show ?thesis
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   368
    using divisor_less_eq_half_nat [of m n] by simp
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   369
qed
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   370
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   371
instance int :: check_prime_by_range
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   372
  apply standard
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   373
  apply (auto simp add: prime_int_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   374
  apply (smt (verit) int_div_less_self)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   375
  apply (rule ccontr)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   376
  apply (auto simp add: neq_iff zdvd_not_zless)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   377
  apply (metis div_by_0 dvd_div_eq_0_iff less_le_not_le one_dvd order_le_less
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   378
      zdvd_not_zless)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   379
  apply (metis atLeastAtMost_iff divisor_less_eq_half_int dvd_div_eq_0_iff
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   380
      int_one_le_iff_zero_less nle_le one_add_one pos_imp_zdiv_nonneg_iff zdiv_eq_0_iff
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   381
      zless_imp_add1_zle)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   382
  done
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   383
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   384
lemma prime_nat_numeral_eq [simp]: \<comment> \<open>TODO Sieve Of Erathosthenes might speed this up\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   385
  "prime (numeral m :: nat) \<longleftrightarrow>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   386
    (1::nat) < numeral m \<and>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   387
    (\<forall>n::nat \<in> set [2..<Suc (numeral m div 2)]. \<not> n dvd numeral m)"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   388
  using prime_iff [of \<open>numeral m :: nat\<close>]
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   389
  by (simp only: set_upt atLeastLessThanSuc_atLeastAtMost)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   390
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   391
context check_prime_by_range
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   392
begin
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   393
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   394
definition check_divisors :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   395
  where \<open>check_divisors l u a \<longleftrightarrow> (\<forall>d\<in>{l..u}. \<not> d dvd a)\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   396
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   397
lemma check_divisors_rec [code]:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   398
  \<open>check_divisors l u a \<longleftrightarrow> u < l \<or> (\<not> l dvd a \<and> check_divisors (l + 1) u a)\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   399
  apply (auto simp add: check_divisors_def not_less)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   400
  apply (metis local.add_increasing2 local.atLeastAtMost_iff local.linear local.order_eq_iff
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   401
      local.zero_le_one)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   402
  subgoal for d
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   403
    apply (cases \<open>l + 1 \<le> d\<close>)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   404
     apply (auto simp add: not_le)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   405
    apply (metis local.dual_order.antisym local.less_eq_iff_succ_less)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   406
    done
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   407
  done
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   408
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   409
lemma prime_eq_check_divisors [code]:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   410
  \<open>prime a \<longleftrightarrow> a > 1 \<and> check_divisors 2 (a div 2) a\<close>
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   411
  by (simp add: check_divisors_def prime_iff)
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   412
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   413
end
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   414
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   415
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   416
subsection \<open>Largest exponent of a prime factor\<close>
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   417
83359
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   418
lemma prime_factor_nat:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   419
  "n \<noteq> (1::nat) \<Longrightarrow> \<exists>p. prime p \<and> p dvd n"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   420
  using prime_divisor_exists[of n]
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   421
  by (cases "n = 0") (auto intro: exI[of _ "2::nat"])
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   422
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   423
lemma prime_factor_int:
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   424
  fixes k :: int
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   425
  assumes "\<bar>k\<bar> \<noteq> 1"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   426
  obtains p where "prime p" "p dvd k"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   427
proof (cases "k = 0")
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   428
  case True
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   429
  then have "prime (2::int)" and "2 dvd k"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   430
    by simp_all
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   431
  with that show thesis
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   432
    by blast
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   433
next
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   434
  case False
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   435
  with assms prime_divisor_exists [of k] obtain p where "prime p" "p dvd k"
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   436
    by auto
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   437
  with that show thesis
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   438
    by blast
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   439
qed
518a1464f1ac more efficient naive prime test
haftmann
parents: 83358
diff changeset
   440
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   441
text\<open>Possibly duplicates other material, but avoid the complexities of multisets.\<close>
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   442
  
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   443
lemma prime_power_cancel_less:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   444
  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and less: "k < k'" and "\<not> p dvd m"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   445
  shows False
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   446
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   447
  obtain l where l: "k' = k + l" and "l > 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   448
    using less less_imp_add_positive by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   449
  have "m = m * (p ^ k) div (p ^ k)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   450
    using \<open>prime p\<close> by simp
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   451
  also have "\<dots> = m' * (p ^ k') div (p ^ k)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   452
    using eq by simp
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   453
  also have "\<dots> = m' * (p ^ l) * (p ^ k) div (p ^ k)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   454
    by (simp add: l mult.commute mult.left_commute power_add)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   455
  also have "... = m' * (p ^ l)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   456
    using \<open>prime p\<close> by simp
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   457
  finally have "p dvd m"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   458
    using \<open>l > 0\<close> by simp
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   459
  with assms show False
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   460
    by simp
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   461
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   462
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   463
lemma prime_power_cancel:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   464
  assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and "\<not> p dvd m" "\<not> p dvd m'"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   465
  shows "k = k'"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   466
  using prime_power_cancel_less [OF \<open>prime p\<close>] assms
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   467
  by (metis linorder_neqE_nat)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   468
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   469
lemma prime_power_cancel2:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   470
  assumes "prime p" "m * (p ^ k) = m' * (p ^ k')" "\<not> p dvd m" "\<not> p dvd m'"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   471
  obtains "m = m'" "k = k'"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   472
  using prime_power_cancel [OF assms] assms by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   473
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   474
lemma prime_power_canonical:
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   475
  fixes m :: nat
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   476
  assumes "prime p" "m > 0"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   477
  shows "\<exists>k n. \<not> p dvd n \<and> m = n * p ^ k"
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   478
using \<open>m > 0\<close>
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   479
proof (induction m rule: less_induct)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   480
  case (less m)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   481
  show ?case
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   482
  proof (cases "p dvd m")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   483
    case True
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   484
    then obtain m' where m': "m = p * m'"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   485
      using dvdE by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   486
    with \<open>prime p\<close> have "0 < m'" "m' < m"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   487
      using less.prems prime_nat_iff by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   488
    with m' less show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   489
      by (metis power_Suc mult.left_commute)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   490
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   491
    case False
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   492
    then show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   493
      by (metis mult.right_neutral power_0)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   494
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   495
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64590
diff changeset
   496
53598
2bd8d459ebae remove unneeded assumption from prime_dvd_power lemmas;
huffman
parents: 47108
diff changeset
   497
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   498
subsection \<open>Infinitely many primes\<close>
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   499
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   500
lemma next_prime_bound: "\<exists>p::nat. prime p \<and> n < p \<and> p \<le> fact n + 1"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   501
proof-
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   502
  have f1: "fact n + 1 \<noteq> (1::nat)" using fact_ge_1 [of n, where 'a=nat] by arith
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   503
  from prime_factor_nat [OF f1]
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   504
  obtain p :: nat where "prime p" and "p dvd fact n + 1" by auto
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   505
  then have "p \<le> fact n + 1" apply (intro dvd_imp_le) apply auto done
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   506
  { assume "p \<le> n"
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   507
    from \<open>prime p\<close> have "p \<ge> 1"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   508
      by (cases p, simp_all)
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   509
    with \<open>p <= n\<close> have "p dvd fact n"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59669
diff changeset
   510
      by (intro dvd_fact)
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   511
    with \<open>p dvd fact n + 1\<close> have "p dvd fact n + 1 - fact n"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   512
      by (rule dvd_diff_nat)
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   513
    then have "p dvd 1" by simp
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   514
    then have "p <= 1" by auto
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 60804
diff changeset
   515
    moreover from \<open>prime p\<close> have "p > 1"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   516
      using prime_nat_iff by blast
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   517
    ultimately have False by auto}
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   518
  then have "n < p" by presburger
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   519
  with \<open>prime p\<close> and \<open>p <= fact n + 1\<close> show ?thesis by auto
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   520
qed
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   521
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   522
lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   523
  using next_prime_bound by auto
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   524
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   525
lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   526
proof
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   527
  assume "finite {(p::nat). prime p}"
67091
1393c2340eec more symbols;
wenzelm
parents: 67051
diff changeset
   528
  with Max_ge have "(\<exists>b. (\<forall>x \<in> {(p::nat). prime p}. x \<le> b))"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   529
    by auto
67091
1393c2340eec more symbols;
wenzelm
parents: 67051
diff changeset
   530
  then obtain b where "\<forall>(x::nat). prime x \<longrightarrow> x \<le> b"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   531
    by auto
44872
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   532
  with bigger_prime [of b] show False
a98ef45122f3 misc tuning;
wenzelm
parents: 44821
diff changeset
   533
    by auto
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   534
qed
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31952
diff changeset
   535
67117
19f627407264 overhauling of primes
haftmann
parents: 67091
diff changeset
   536
subsection \<open>Powers of Primes\<close>
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   537
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   538
text\<open>Versions for type nat only\<close>
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   539
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   540
lemma prime_product:
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   541
  fixes p::nat
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   542
  assumes "prime (p * q)"
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   543
  shows "p = 1 \<or> q = 1"
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   544
proof -
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   545
  from assms have
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   546
    "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   547
    unfolding prime_nat_iff by auto
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   548
  from \<open>1 < p * q\<close> have "p \<noteq> 0" by (cases p) auto
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   549
  then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   550
  have "p dvd p * q" by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   551
  then have "p = 1 \<or> p = p * q" by (rule P)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   552
  then show ?thesis by (simp add: Q)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   553
qed
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   554
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   555
(* TODO: Generalise? *)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   556
lemma prime_power_mult_nat:
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   557
  fixes p :: nat
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   558
  assumes p: "prime p" and xy: "x * y = p ^ k"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   559
  shows "\<exists>i j. x = p ^ i \<and> y = p^ j"
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   560
using xy
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   561
proof(induct k arbitrary: x y)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   562
  case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   563
next
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   564
  case (Suc k x y)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   565
  from Suc.prems have pxy: "p dvd x*y" by auto
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   566
  from prime_dvd_multD [OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   567
  from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   568
  {assume px: "p dvd x"
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   569
    then obtain d where d: "x = p*d" unfolding dvd_def by blast
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   570
    from Suc.prems d  have "p*d*y = p^Suc k" by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   571
    hence th: "d*y = p^k" using p0 by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   572
    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   573
    with d have "x = p^Suc i" by simp
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   574
    with ij(2) have ?case by blast}
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   575
  moreover
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   576
  {assume px: "p dvd y"
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   577
    then obtain d where d: "y = p*d" unfolding dvd_def by blast
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55337
diff changeset
   578
    from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult.commute)
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   579
    hence th: "d*x = p^k" using p0 by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   580
    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   581
    with d have "y = p^Suc i" by simp
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   582
    with ij(2) have ?case by blast}
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   583
  ultimately show ?case  using pxyc by blast
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   584
qed
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   585
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   586
lemma prime_power_exp_nat:
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   587
  fixes p::nat
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   588
  assumes p: "prime p" and n: "n \<noteq> 0"
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   589
    and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   590
  using n xn
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   591
proof(induct n arbitrary: k)
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   592
  case 0 thus ?case by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   593
next
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   594
  case (Suc n k) hence th: "x*x^n = p^k" by simp
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   595
  {assume "n = 0" with Suc have ?case by simp (rule exI[where x="k"], simp)}
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   596
  moreover
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   597
  {assume n: "n \<noteq> 0"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   598
    from prime_power_mult_nat[OF p th]
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   599
    obtain i j where ij: "x = p^i" "x^n = p^j"by blast
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   600
    from Suc.hyps[OF n ij(2)] have ?case .}
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   601
  ultimately show ?case by blast
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   602
qed
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   603
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   604
lemma divides_primepow_nat:
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   605
  fixes p :: nat
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   606
  assumes p: "prime p"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   607
  shows "d dvd p ^ k \<longleftrightarrow> (\<exists>i\<le>k. d = p ^ i)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   608
  using assms divides_primepow [of p d k] by (auto intro: le_imp_power_dvd)
55215
b6c926e67350 Restoring some proofs from the equivalent file in Old_Number_Theory.
paulson <lp15@cam.ac.uk>
parents: 55130
diff changeset
   609
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   610
lemma gcd_prime_int:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   611
  assumes "prime (p :: int)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   612
  shows   "gcd p k = (if p dvd k then p else 1)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   613
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   614
  have "p \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   615
    using assms prime_ge_0_int by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   616
  show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   617
  proof (cases "p dvd k")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   618
    case True
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   619
    thus ?thesis using assms \<open>p \<ge> 0\<close> by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   620
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   621
    case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   622
    hence "coprime p k"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   623
      using assms by (simp add: prime_imp_coprime)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   624
    with False show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   625
      by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   626
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   627
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 80084
diff changeset
   628
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   629
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   630
subsection \<open>Chinese Remainder Theorem Variants\<close>
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   631
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   632
lemma bezout_gcd_nat:
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   633
  fixes a::nat shows "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   634
  using bezout_nat[of a b]
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62410
diff changeset
   635
by (metis bezout_nat diff_add_inverse gcd_add_mult gcd.commute
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   636
  gcd_nat.right_neutral mult_0)
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   637
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   638
lemma gcd_bezout_sum_nat:
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   639
  fixes a::nat
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   640
  assumes "a * x + b * y = d"
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   641
  shows "gcd a b dvd d"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   642
proof-
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   643
  let ?g = "gcd a b"
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   644
    have dv: "?g dvd a*x" "?g dvd b * y"
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   645
      by simp_all
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   646
    from dvd_add[OF dv] assms
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   647
    show ?thesis by auto
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   648
qed
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   649
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   650
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   651
text \<open>A binary form of the Chinese Remainder Theorem.\<close>
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   652
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   653
(* TODO: Generalise? *)
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   654
lemma chinese_remainder:
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   655
  fixes a::nat  assumes ab: "coprime a b" and a: "a \<noteq> 0" and b: "b \<noteq> 0"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   656
  shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   657
proof-
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   658
  from bezout_add_strong_nat[OF a, of b] bezout_add_strong_nat[OF b, of a]
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   659
  obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   660
    and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   661
  then have d12: "d1 = 1" "d2 = 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   662
    using ab coprime_common_divisor_nat [of a b] by blast+
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   663
  let ?x = "v * a * x1 + u * b * x2"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   664
  let ?q1 = "v * x1 + u * y2"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   665
  let ?q2 = "v * y1 + u * x2"
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   666
  from dxy2(3)[simplified d12] dxy1(3)[simplified d12]
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   667
  have "?x = u + ?q1 * a" "?x = v + ?q2 * b"
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
   668
    by algebra+
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   669
  thus ?thesis by blast
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   670
qed
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   671
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 59730
diff changeset
   672
text \<open>Primality\<close>
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   673
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   674
lemma coprime_bezout_strong:
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   675
  fixes a::nat assumes "coprime a b"  "b \<noteq> 1"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   676
  shows "\<exists>x y. a * x = b * y + 1"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   677
  by (metis add.commute add.right_neutral assms(1) assms(2) chinese_remainder coprime_1_left coprime_1_right coprime_crossproduct_nat mult.commute mult.right_neutral mult_cancel_left)
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   678
59669
de7792ea4090 renaming HOL/Fact.thy -> Binomial.thy
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
   679
lemma bezout_prime:
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   680
  assumes p: "prime p" and pa: "\<not> p dvd a"
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   681
  shows "\<exists>x y. a*x = Suc (p*y)"
62349
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   682
proof -
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   683
  have ap: "coprime a p"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66837
diff changeset
   684
    using coprime_commute p pa prime_imp_coprime by auto
62349
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   685
  moreover from p have "p \<noteq> 1" by auto
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   686
  ultimately have "\<exists>x y. a * x = p * y + 1"
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   687
    by (rule coprime_bezout_strong)
7c23469b5118 cleansed junk-producing interpretations for gcd/lcm on nat altogether
haftmann
parents: 62348
diff changeset
   688
  then show ?thesis by simp    
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   689
qed
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   690
(* END TODO *)
55238
7ddb889e23bd Added material from Old_Number_Theory related to the Chinese Remainder Theorem
paulson <lp15@cam.ac.uk>
parents: 55215
diff changeset
   691
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   692
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   693
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   694
subsection \<open>Multiplicity and primality for natural numbers and integers\<close>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   695
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   696
lemma prime_factors_gt_0_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   697
  "p \<in> prime_factors x \<Longrightarrow> p > (0::nat)"
63905
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   698
  by (simp add: in_prime_factors_imp_prime prime_gt_0_nat)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   699
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   700
lemma prime_factors_gt_0_int:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   701
  "p \<in> prime_factors x \<Longrightarrow> p > (0::int)"
63905
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   702
  by (simp add: in_prime_factors_imp_prime prime_gt_0_int)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   703
63905
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   704
lemma prime_factors_ge_0_int [elim]: (* FIXME !? *)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   705
  fixes n :: int
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   706
  shows "p \<in> prime_factors n \<Longrightarrow> p \<ge> 0"
63905
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   707
  by (drule prime_factors_gt_0_int) simp
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   708
  
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   709
lemma prod_mset_prime_factorization_int:
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   710
  fixes n :: int
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   711
  assumes "n > 0"
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   712
  shows   "prod_mset (prime_factorization n) = n"
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   713
  using assms by (simp add: prod_mset_prime_factorization)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   714
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   715
lemma prime_factorization_exists_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   716
  "n > 0 \<Longrightarrow> (\<exists>M. (\<forall>p::nat \<in> set_mset M. prime p) \<and> n = (\<Prod>i \<in># M. i))"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67117
diff changeset
   717
  using prime_factorization_exists[of n] by auto
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   718
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   719
lemma prod_mset_prime_factorization_nat [simp]: 
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   720
  "(n::nat) > 0 \<Longrightarrow> prod_mset (prime_factorization n) = n"
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   721
  by (subst prod_mset_prime_factorization) simp_all
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   722
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   723
lemma prime_factorization_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   724
    "n > (0::nat) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   725
  by (simp add: prod_prime_factors)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   726
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   727
lemma prime_factorization_int:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   728
    "n > (0::int) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   729
  by (simp add: prod_prime_factors)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   730
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   731
lemma prime_factorization_unique_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   732
  fixes f :: "nat \<Rightarrow> _"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   733
  assumes S_eq: "S = {p. 0 < f p}"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   734
    and "finite S"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   735
    and S: "\<forall>p\<in>S. prime p" "n = (\<Prod>p\<in>S. p ^ f p)"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   736
  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   737
  using assms by (intro prime_factorization_unique'') auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   738
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   739
lemma prime_factorization_unique_int:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   740
  fixes f :: "int \<Rightarrow> _"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   741
  assumes S_eq: "S = {p. 0 < f p}"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   742
    and "finite S"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   743
    and S: "\<forall>p\<in>S. prime p" "abs n = (\<Prod>p\<in>S. p ^ f p)"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   744
  shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   745
  using assms by (intro prime_factorization_unique'') auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   746
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   747
lemma prime_factors_characterization_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   748
  "S = {p. 0 < f (p::nat)} \<Longrightarrow>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   749
    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   750
  by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric])
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   751
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   752
lemma prime_factors_characterization'_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   753
  "finite {p. 0 < f (p::nat)} \<Longrightarrow>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   754
    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   755
      prime_factors (\<Prod>p | 0 < f p. p ^ f p) = {p. 0 < f p}"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   756
  by (rule prime_factors_characterization_nat) auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   757
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   758
lemma prime_factors_characterization_int:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   759
  "S = {p. 0 < f (p::int)} \<Longrightarrow> finite S \<Longrightarrow>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   760
    \<forall>p\<in>S. prime p \<Longrightarrow> abs n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   761
  by (rule prime_factorization_unique_int [THEN conjunct1, symmetric])
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   762
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   763
(* TODO Move *)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   764
lemma abs_prod: "abs (prod f A :: 'a :: linordered_idom) = prod (\<lambda>x. abs (f x)) A"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   765
  by (cases "finite A", induction A rule: finite_induct) (simp_all add: abs_mult)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   766
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   767
lemma primes_characterization'_int [rule_format]:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   768
  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow> \<forall>p. 0 < f p \<longrightarrow> prime p \<Longrightarrow>
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   769
      prime_factors (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = {p. p \<ge> 0 \<and> 0 < f p}"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   770
  by (rule prime_factors_characterization_int) (auto simp: abs_prod prime_ge_0_int)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   771
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   772
lemma multiplicity_characterization_nat:
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   773
  "S = {p. 0 < f (p::nat)} \<Longrightarrow> finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow>
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   774
    n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   775
  by (frule prime_factorization_unique_nat [of S f n, THEN conjunct2, rule_format, symmetric]) auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   776
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   777
lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   778
    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<longrightarrow> prime p \<longrightarrow>
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   779
      multiplicity p (\<Prod>p | 0 < f p. p ^ f p) = f p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   780
  by (intro impI, rule multiplicity_characterization_nat) auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   781
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   782
lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   783
    finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   784
  by (frule prime_factorization_unique_int [of S f n, THEN conjunct2, rule_format, symmetric]) 
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   785
     (auto simp: abs_prod power_abs prime_ge_0_int intro!: prod.cong)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   786
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   787
lemma multiplicity_characterization'_int [rule_format]:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   788
  "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow>
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   789
    (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> prime p \<Longrightarrow>
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   790
      multiplicity p (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = f p"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   791
  by (rule multiplicity_characterization_int) (auto simp: prime_ge_0_int)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   792
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   793
lemma multiplicity_one_nat [simp]: "multiplicity p (Suc 0) = 0"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   794
  unfolding One_nat_def [symmetric] by (rule multiplicity_one)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   795
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   796
lemma multiplicity_eq_nat:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   797
  fixes x and y::nat
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   798
  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   799
  shows "x = y"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   800
  using multiplicity_eq_imp_eq[of x y] assms by simp
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   801
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   802
lemma multiplicity_eq_int:
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   803
  fixes x y :: int
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   804
  assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   805
  shows "x = y"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   806
  using multiplicity_eq_imp_eq[of x y] assms by simp
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   807
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   808
lemma multiplicity_prod_prime_powers:
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
   809
  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> prime x" "prime p"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   810
  shows   "multiplicity p (\<Prod>p \<in> S. p ^ f p) = (if p \<in> S then f p else 0)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   811
proof -
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   812
  define g where "g = (\<lambda>x. if x \<in> S then f x else 0)"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   813
  define A where "A = Abs_multiset g"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   814
  have "{x. g x > 0} \<subseteq> S" by (auto simp: g_def)
73270
e2d03448d5b5 get rid of traditional predicate
haftmann
parents: 73047
diff changeset
   815
  from finite_subset[OF this assms(1)] have [simp]: "finite {x. 0 < g x}"
e2d03448d5b5 get rid of traditional predicate
haftmann
parents: 73047
diff changeset
   816
    by simp
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   817
  from assms have count_A: "count A x = g x" for x unfolding A_def
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   818
    by simp
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   819
  have set_mset_A: "set_mset A = {x\<in>S. f x > 0}"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   820
    unfolding set_mset_def count_A by (auto simp: g_def)
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   821
  with assms have prime: "prime x" if "x \<in># A" for x using that by auto
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   822
  from set_mset_A assms have "(\<Prod>p \<in> S. p ^ f p) = (\<Prod>p \<in> S. p ^ g p) "
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   823
    by (intro prod.cong) (auto simp: g_def)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   824
  also from set_mset_A assms have "\<dots> = (\<Prod>p \<in> set_mset A. p ^ g p)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   825
    by (intro prod.mono_neutral_right) (auto simp: g_def set_mset_A)
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   826
  also have "\<dots> = prod_mset A"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   827
    by (auto simp: prod_mset_multiplicity count_A set_mset_A intro!: prod.cong)
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   828
  also from assms have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   829
    by (subst prime_elem_multiplicity_prod_mset_distrib) (auto dest: prime)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   830
  also from assms have "image_mset (multiplicity p) A = image_mset (\<lambda>x. if x = p then 1 else 0) A"
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   831
    by (intro image_mset_cong) (auto simp: prime_multiplicity_other dest: prime)
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63766
diff changeset
   832
  also have "sum_mset \<dots> = (if p \<in> S then f p else 0)" by (simp add: sum_mset_delta count_A g_def)
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   833
  finally show ?thesis .
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   834
qed
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
   835
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   836
lemma prime_factorization_prod_mset:
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   837
  assumes "0 \<notin># A"
73047
ab9e27da0e85 HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents: 69597
diff changeset
   838
  shows "prime_factorization (prod_mset A) = \<Sum>\<^sub>#(image_mset prime_factorization A)"
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   839
  using assms by (induct A) (auto simp add: prime_factorization_mult)
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   840
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   841
lemma prime_factors_prod:
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   842
  assumes "finite A" and "0 \<notin> f ` A"
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 68623
diff changeset
   843
  shows "prime_factors (prod f A) = \<Union>((prime_factors \<circ> f) ` A)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   844
  using assms by (simp add: prod_unfold_prod_mset prime_factorization_prod_mset)
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   845
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   846
lemma prime_factors_fact:
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   847
  "prime_factors (fact n) = {p \<in> {2..n}. prime p}" (is "?M = ?N")
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   848
proof (rule set_eqI)
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   849
  fix p
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   850
  { fix m :: nat
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   851
    assume "p \<in> prime_factors m"
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   852
    then have "prime p" and "p dvd m" by auto
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   853
    moreover assume "m > 0" 
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   854
    ultimately have "2 \<le> p" and "p \<le> m"
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   855
      by (auto intro: prime_ge_2_nat dest: dvd_imp_le)
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   856
    moreover assume "m \<le> n"
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   857
    ultimately have "2 \<le> p" and "p \<le> n"
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   858
      by (auto intro: order_trans)
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   859
  } note * = this
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   860
  show "p \<in> ?M \<longleftrightarrow> p \<in> ?N"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 63905
diff changeset
   861
    by (auto simp add: fact_prod prime_factors_prod Suc_le_eq dest!: prime_prime_factors intro: *)
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   862
qed
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   863
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63635
diff changeset
   864
lemma prime_dvd_fact_iff:
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63635
diff changeset
   865
  assumes "prime p"
63904
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   866
  shows "p dvd fact n \<longleftrightarrow> p \<le> n"
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   867
  using assms
b8482e12a2a8 more lemmas
haftmann
parents: 63830
diff changeset
   868
  by (auto simp add: prime_factorization_subset_iff_dvd [symmetric]
63905
1c3dcb5fe6cb prefer abbreviation for trivial set conversion
haftmann
parents: 63904
diff changeset
   869
    prime_factorization_prime prime_factors_fact prime_ge_2_nat)
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63635
diff changeset
   870
79544
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   871
lemma dvd_choose_prime:
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   872
  assumes kn: "k < n" and k: "k \<noteq> 0" and n: "n \<noteq> 0" and prime_n: "prime n"
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   873
  shows "n dvd (n choose k)"
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   874
proof -
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   875
  have "n dvd (fact n)" by (simp add: fact_num_eq_if n)
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   876
  moreover have "\<not> n dvd (fact k * fact (n-k))"
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   877
    by (metis prime_dvd_fact_iff prime_dvd_mult_iff assms neq0_conv diff_less linorder_not_less)
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   878
  moreover have "(fact n::nat) = fact k * fact (n-k) * (n choose k)"
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   879
    using binomial_fact_lemma kn by auto
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   880
  ultimately show ?thesis using prime_n
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   881
    by (auto simp add: prime_dvd_mult_iff)
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   882
qed
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
   883
80084
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   884
lemma (in ring_1) minus_power_prime_CHAR:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   885
  assumes "p = CHAR('a)" "prime p"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   886
  shows "(-x :: 'a) ^ p = -(x ^ p)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   887
proof (cases "p = 2")
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   888
  case False
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   889
  have "prime p"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   890
    using assms by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   891
  hence "odd p"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   892
    using prime_imp_coprime assms False coprime_right_2_iff_odd gcd_nat.strict_iff_not by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   893
  thus ?thesis
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   894
    by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   895
qed (use assms in \<open>auto simp: uminus_CHAR_2\<close>)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   896
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   897
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   898
subsection \<open>Rings and fields with prime characteristic\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   899
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   900
text \<open>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   901
  We introduce some type classes for rings and fields with prime characteristic.
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   902
\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   903
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   904
class semiring_prime_char = semiring_1 +
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   905
  assumes prime_char_aux: "\<exists>n. prime n \<and> of_nat n = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   906
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   907
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   908
lemma CHAR_pos [intro, simp]: "CHAR('a) > 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   909
  using local.CHAR_pos_iff local.prime_char_aux prime_gt_0_nat by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   910
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   911
lemma CHAR_nonzero [simp]: "CHAR('a) \<noteq> 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   912
  using CHAR_pos by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   913
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   914
lemma CHAR_prime [intro, simp]: "prime CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   915
  by (metis (mono_tags, lifting) gcd_nat.order_iff_strict local.of_nat_1 local.of_nat_eq_0_iff_char_dvd
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   916
        local.one_neq_zero local.prime_char_aux prime_nat_iff)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   917
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   918
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   919
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   920
lemma semiring_prime_charI [intro?]:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   921
  "prime CHAR('a :: semiring_1) \<Longrightarrow> OFCLASS('a, semiring_prime_char_class)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   922
  by standard auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   923
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   924
lemma idom_prime_charI [intro?]:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   925
  assumes "CHAR('a :: idom) > 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   926
  shows   "OFCLASS('a, semiring_prime_char_class)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   927
proof
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   928
  show "prime CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   929
    using assms prime_CHAR_semidom by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   930
qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   931
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   932
class comm_semiring_prime_char = comm_semiring_1 + semiring_prime_char
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   933
class comm_ring_prime_char = comm_ring_1 + semiring_prime_char
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   934
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   935
subclass comm_semiring_prime_char ..
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   936
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   937
class idom_prime_char = idom + semiring_prime_char
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   938
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   939
subclass comm_ring_prime_char ..
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   940
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   941
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   942
class field_prime_char = field +
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   943
  assumes pos_char_exists: "\<exists>n>0. of_nat n = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   944
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   945
subclass idom_prime_char
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   946
  apply standard
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   947
  using pos_char_exists local.CHAR_pos_iff local.of_nat_CHAR local.prime_CHAR_semidom by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   948
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   949
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   950
lemma field_prime_charI [intro?]:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   951
  "n > 0 \<Longrightarrow> of_nat n = (0 :: 'a :: field) \<Longrightarrow> OFCLASS('a, field_prime_char_class)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   952
  by standard auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   953
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   954
lemma field_prime_charI' [intro?]:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   955
  "CHAR('a :: field) > 0 \<Longrightarrow> OFCLASS('a, field_prime_char_class)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   956
  by standard auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   957
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   958
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   959
subsection \<open>Finite fields\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   960
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   961
class finite_field = field_prime_char + finite
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   962
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   963
lemma finite_fieldI [intro?]:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   964
  assumes "finite (UNIV :: 'a :: field set)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   965
  shows   "OFCLASS('a, finite_field_class)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   966
proof standard
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   967
  show "\<exists>n>0. of_nat n = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   968
    using assms prime_CHAR_semidom[where ?'a = 'a] finite_imp_CHAR_pos[OF assms]
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   969
    by (intro exI[of _ "CHAR('a)"]) auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   970
qed fact+
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   971
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   972
text \<open>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   973
  On a finite field with \<open>n\<close> elements, taking the \<open>n\<close>-th power of an element
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   974
  is the identity. This is an obvious consequence of the fact that the multiplicative group of
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   975
  the field is a finite group of order \<open>n - 1\<close>, so \<open>x^n = 1\<close> for any non-zero \<open>x\<close>.
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   976
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   977
  Note that this result is sharp in the sense that the multiplicative group of a
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   978
  finite field is cyclic, i.e.\ it contains an element of order \<open>n - 1\<close>.
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   979
  (We don't prove this here.)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   980
\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   981
lemma finite_field_power_card_eq_same:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   982
  fixes x :: "'a :: finite_field"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   983
  shows   "x ^ card (UNIV :: 'a set) = x"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   984
proof (cases "x = 0")
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   985
  case False
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   986
  have "x * (\<Prod>y\<in>UNIV-{0}. x * y) = x * x ^ (card (UNIV :: 'a set) - 1) * \<Prod>(UNIV-{0})"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   987
    by (simp add: prod.distrib mult_ac)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   988
  also have "x * x ^ (card (UNIV :: 'a set) - 1) = x ^ Suc (card (UNIV :: 'a set) - 1)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   989
    by (subst power_Suc) auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   990
  also have "Suc (card (UNIV :: 'a set) - 1) = card (UNIV :: 'a set)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   991
    using finite_UNIV_card_ge_0[where ?'a = 'a] by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   992
  also have "(\<Prod>y\<in>UNIV-{0}. x * y) = (\<Prod>y\<in>UNIV-{0}. y)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   993
    by (rule prod.reindex_bij_witness[of _ "\<lambda>y. y / x" "\<lambda>y. x * y"]) (use False in auto)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   994
  finally show ?thesis
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   995
    by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   996
qed (use finite_UNIV_card_ge_0[where ?'a = 'a] in auto)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   997
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   998
lemma finite_field_power_card_power_eq_same:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
   999
  fixes x :: "'a :: finite_field"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1000
  assumes "m = card (UNIV :: 'a set) ^ n"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1001
  shows   "x ^ m = x"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1002
  unfolding assms
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1003
  by (induction n) (simp_all add: finite_field_power_card_eq_same power_mult)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1004
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1005
class enum_finite_field = finite_field +
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1006
  fixes enum_finite_field :: "nat \<Rightarrow> 'a"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1007
  assumes enum_finite_field: "enum_finite_field ` {..<card (UNIV :: 'a set)} = UNIV"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1008
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1009
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1010
lemma inj_on_enum_finite_field: "inj_on enum_finite_field {..<card (UNIV :: 'a set)}"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1011
  using enum_finite_field by (simp add: eq_card_imp_inj_on)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1012
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1013
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1014
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1015
text \<open>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1016
  To get rid of the pending sort hypotheses, we prove that the field with 2 elements is indeed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1017
  a finite field.
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1018
\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1019
typedef gf2 = "{0, 1 :: nat}"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1020
  by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1021
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1022
setup_lifting type_definition_gf2
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1023
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1024
instantiation gf2 :: field
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1025
begin
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1026
lift_definition zero_gf2 :: "gf2" is "0" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1027
lift_definition one_gf2 :: "gf2" is "1" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1028
lift_definition uminus_gf2 :: "gf2 \<Rightarrow> gf2" is "\<lambda>x. x" .
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1029
lift_definition plus_gf2 :: "gf2 \<Rightarrow> gf2 \<Rightarrow> gf2" is "\<lambda>x y. if x = y then 0 else 1" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1030
lift_definition minus_gf2 :: "gf2 \<Rightarrow> gf2 \<Rightarrow> gf2" is "\<lambda>x y. if x = y then 0 else 1" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1031
lift_definition times_gf2 :: "gf2 \<Rightarrow> gf2 \<Rightarrow> gf2" is "\<lambda>x y. x * y" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1032
lift_definition inverse_gf2 :: "gf2 \<Rightarrow> gf2" is "\<lambda>x. x" .
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1033
lift_definition divide_gf2 :: "gf2 \<Rightarrow> gf2 \<Rightarrow> gf2" is "\<lambda>x y. x * y" by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1034
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1035
instance
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1036
  by standard (transfer; fastforce)+
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1037
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1038
end
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1039
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1040
instance gf2 :: finite_field
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1041
proof
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1042
  interpret type_definition Rep_gf2 Abs_gf2 "{0, 1 :: nat}"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1043
    by (rule type_definition_gf2)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1044
  show "finite (UNIV :: gf2 set)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1045
    by (metis Abs_image finite.emptyI finite.insertI finite_imageI)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1046
qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1047
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1048
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1049
subsection \<open>The Freshman's Dream in rings of prime characteristic\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1050
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1051
lemma (in comm_semiring_1) freshmans_dream:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1052
  fixes x y :: 'a and n :: nat
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1053
  assumes "prime CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1054
  assumes n_def: "n = CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1055
  shows   "(x + y) ^ n = x ^ n + y ^ n"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1056
proof -
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1057
  interpret comm_semiring_prime_char
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1058
    by standard (auto intro!: exI[of _ "CHAR('a)"] assms)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1059
  have "n > 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1060
    unfolding n_def by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1061
  have "(x + y) ^ n = (\<Sum>k\<le>n. of_nat (n choose k) * x ^ k * y ^ (n - k))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1062
    by (rule binomial_ring)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1063
  also have "\<dots> = (\<Sum>k\<in>{0,n}. of_nat (n choose k) * x ^ k * y ^ (n - k))"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1064
  proof (intro sum.mono_neutral_right ballI)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1065
    fix k assume "k \<in> {..n} - {0, n}"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1066
    hence k: "k > 0" "k < n"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1067
      by auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1068
    have "CHAR('a) dvd (n choose k)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1069
      unfolding n_def
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1070
      by (rule dvd_choose_prime) (use k in \<open>auto simp: n_def\<close>)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1071
    hence "of_nat (n choose k) = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1072
      using of_nat_eq_0_iff_char_dvd by blast
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1073
    thus "of_nat (n choose k) * x ^ k * y ^ (n - k) = 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1074
      by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1075
  qed auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1076
  finally show ?thesis
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1077
    using \<open>n > 0\<close> by (simp add: add_ac)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1078
qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1079
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1080
lemma (in comm_semiring_1) freshmans_dream':
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1081
  assumes [simp]: "prime CHAR('a)" and "m = CHAR('a) ^ n"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1082
  shows "(x + y :: 'a) ^ m = x ^ m + y ^ m"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1083
  unfolding assms(2)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1084
proof (induction n)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1085
  case (Suc n)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1086
  have "(x + y) ^ (CHAR('a) ^ n * CHAR('a)) = ((x + y) ^ (CHAR('a) ^ n)) ^ CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1087
    by (rule power_mult)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1088
  thus ?case
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1089
    by (simp add: Suc.IH freshmans_dream Groups.mult_ac flip: power_mult)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1090
qed auto
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1091
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1092
lemma (in comm_semiring_1) freshmans_dream_sum:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1093
  fixes f :: "'b \<Rightarrow> 'a"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1094
  assumes "prime CHAR('a)" and "n = CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1095
  shows "sum f A ^ n = sum (\<lambda>i. f i ^ n) A"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1096
  using assms
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1097
  by (induct A rule: infinite_finite_induct)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1098
     (auto simp add: power_0_left freshmans_dream)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1099
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1100
lemma (in comm_semiring_1) freshmans_dream_sum':
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1101
  fixes f :: "'b \<Rightarrow> 'a"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1102
  assumes "prime CHAR('a)" "m = CHAR('a) ^ n"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1103
  shows   "sum f A ^ m = sum (\<lambda>i. f i ^ m) A"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1104
  using assms
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1105
  by (induction A rule: infinite_finite_induct)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1106
     (auto simp: freshmans_dream' power_0_left)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1107
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 79544
diff changeset
  1108
79544
50ee2921da94 A few more new theorems taken from AFP entries
paulson <lp15@cam.ac.uk>
parents: 73270
diff changeset
  1109
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
  1110
(* TODO Legacy names *)
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1111
lemmas prime_imp_coprime_nat = prime_imp_coprime[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1112
lemmas prime_imp_coprime_int = prime_imp_coprime[where ?'a = int]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1113
lemmas prime_dvd_mult_nat = prime_dvd_mult_iff[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1114
lemmas prime_dvd_mult_int = prime_dvd_mult_iff[where ?'a = int]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1115
lemmas prime_dvd_mult_eq_nat = prime_dvd_mult_iff[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1116
lemmas prime_dvd_mult_eq_int = prime_dvd_mult_iff[where ?'a = int]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1117
lemmas prime_dvd_power_nat = prime_dvd_power[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1118
lemmas prime_dvd_power_int = prime_dvd_power[where ?'a = int]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1119
lemmas prime_dvd_power_nat_iff = prime_dvd_power_iff[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1120
lemmas prime_dvd_power_int_iff = prime_dvd_power_iff[where ?'a = int]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1121
lemmas prime_imp_power_coprime_nat = prime_imp_power_coprime[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1122
lemmas prime_imp_power_coprime_int = prime_imp_power_coprime[where ?'a = int]
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
  1123
lemmas primes_coprime_nat = primes_coprime[where ?'a = nat]
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
  1124
lemmas primes_coprime_int = primes_coprime[where ?'a = nat]
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1125
lemmas prime_divprod_pow_nat = prime_elem_divprod_pow[where ?'a = nat]
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63535
diff changeset
  1126
lemmas prime_exp = prime_elem_power_iff[where ?'a = nat]
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62481
diff changeset
  1127
63635
858a225ebb62 Tuned primes
eberlm <eberlm@in.tum.de>
parents: 63633
diff changeset
  1128
end