src/HOL/Number_Theory/Pocklington.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 82690 cccbfa567117
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(*  Title:      HOL/Number_Theory/Pocklington.thy
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
     2
    Author:     Amine Chaieb, Manuel Eberl
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
*)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>Pocklington's Theorem for Primes\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Pocklington
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
imports Residues
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    11
subsection \<open>Lemmas about previously defined terms\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    13
lemma prime_nat_iff'': "prime (p::nat) \<longleftrightarrow> p \<noteq> 0 \<and> p \<noteq> 1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
82690
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    14
proof -
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    15
have \<section>: "\<And>m. \<lbrakk>0 < p; \<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m; m dvd p; m \<noteq> p\<rbrakk>
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    16
         \<Longrightarrow> m = Suc 0"
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    17
  by (metis One_nat_def coprime_absorb_right dvd_1_iff_1 dvd_nat_bounds
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    18
      nless_le)
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    19
  show ?thesis
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    20
    by (auto simp: nat_dvd_not_less prime_imp_coprime_nat prime_nat_iff elim!: \<section>)
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    21
qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    24
proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    26
  then show ?thesis by simp
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    30
subsection \<open>Some basic theorems about solving congruences\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    32
lemma cong_solve:
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    33
  fixes n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    34
  assumes an: "coprime a n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    35
  shows "\<exists>x. [a * x = b] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    36
proof (cases "a = 0")
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    37
  case True
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    38
  with an show ?thesis
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
    39
    by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    40
next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    41
  case False
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    42
  from bezout_add_strong_nat [OF this]
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    43
  obtain d x y where dxy: "d dvd a" "d dvd n" "a * x = n * y + d" by blast
82690
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
    44
  then have d1: "d = 1"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    45
    using assms coprime_common_divisor [of a n d] by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    46
  with dxy(3) have "a * x * b = (n * y + 1) * b"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    47
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    48
  then have "a * (x * b) = n * (y * b) + b"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    49
    by (auto simp: algebra_simps)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    50
  then have "a * (x * b) mod n = (n * (y * b) + b) mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    51
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    52
  then have "a * (x * b) mod n = b mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    53
    by (simp add: mod_add_left_eq)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    54
  then have "[a * (x * b) = b] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
    55
    by (simp only: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    56
  then show ?thesis by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    59
lemma cong_solve_unique:
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    60
  fixes n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    61
  assumes an: "coprime a n" and nz: "n \<noteq> 0"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
  shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    63
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    64
  from cong_solve[OF an] obtain x where x: "[a * x = b] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    65
    by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  let ?x = "x mod n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    68
  from x have *: "[a * ?x = b] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
    69
    by (simp add: cong_def mod_mult_right_eq[of a x n])
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    70
  from mod_less_divisor[ of n x] nz * have Px: "?P ?x" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    71
  have "y = ?x" if Py: "y < n" "[a * y = b] (mod n)" for y
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    72
  proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    73
    from Py(2) * have "[a * y = a * ?x] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
    74
      by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    75
    then have "[y = ?x] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    76
      by (metis an cong_mult_lcancel_nat)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    78
    show ?thesis
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
    79
      by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    80
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  with Px show ?thesis by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
lemma cong_solve_unique_nontrivial:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    85
  fixes p :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    86
  assumes p: "prime p"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    87
    and pa: "coprime p a"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    88
    and x0: "0 < x"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    89
    and xp: "x < p"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    91
proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
  from pa have ap: "coprime a p"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    93
    by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    94
  from x0 xp p have px: "coprime x p"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    95
    by (auto simp add: prime_nat_iff'' ac_simps)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
    by (metis cong_solve_unique neq0_conv p prime_gt_0_nat px)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    98
  have "y \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
    99
  proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   100
    assume "y = 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   101
    with y(2) have "p dvd a"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   102
      using cong_dvd_iff by auto
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   103
    with not_prime_1 p pa show False
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   104
      by (auto simp add: gcd_nat.order_iff)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   105
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   106
  with y show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   107
    by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
lemma cong_unique_inverse_prime:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   111
  fixes p :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   112
  assumes "prime p" and "0 < x" and "x < p"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   114
  by (rule cong_solve_unique_nontrivial) (use assms in simp_all)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
lemma chinese_remainder_coprime_unique:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   117
  fixes a :: nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   119
    and ma: "coprime m a" and nb: "coprime n b"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   121
proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
  from binary_chinese_remainder_unique_nat[OF ab az bz]
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   124
  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" "\<forall>y. ?P y \<longrightarrow> y = x"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   125
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   126
  from ma nb x have "coprime x a" "coprime x b"
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 68157
diff changeset
   127
    using cong_imp_coprime cong_sym by blast+
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  then have "coprime x (a*b)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   129
    by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   130
  with x show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   131
    by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   135
subsection \<open>Lucas's theorem\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
lemma lucas_coprime_lemma:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   138
  fixes n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   139
  assumes m: "m \<noteq> 0" and am: "[a^m = 1] (mod n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  shows "coprime a n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   141
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   142
  consider "n = 1" | "n = 0" | "n > 1" by arith
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   143
  then show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   144
  proof cases
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   145
    case 1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   146
    then show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   147
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   148
    case 2
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   149
    with am m show ?thesis
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   150
      by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   151
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   152
    case 3
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   153
    from m obtain m' where m': "m = Suc m'" by (cases m) blast+
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   154
    have "d = 1" if d: "d dvd a" "d dvd n" for d
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   155
    proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   156
      from am mod_less[OF \<open>n > 1\<close>] have am1: "a^m mod n = 1"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   157
        by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   158
      from dvd_mult2[OF d(1), of "a^m'"] have dam: "d dvd a^m"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   159
        by (simp add: m')
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   160
      from dvd_mod_iff[OF d(2), of "a^m"] dam am1 show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   161
        by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   162
    qed
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   163
    then show ?thesis
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   164
      by (auto intro: coprimeI)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   165
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
lemma lucas_weak:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   169
  fixes n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   170
  assumes n: "n \<ge> 2"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   171
    and an: "[a ^ (n - 1) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   172
    and nm: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  shows "prime n"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   174
proof (rule totient_imp_prime)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   175
  show "totient n = n - 1"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   176
  proof (rule ccontr)
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   177
    have "[a ^ totient n = 1] (mod n)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   178
      by (rule euler_theorem, rule lucas_coprime_lemma [of "n - 1"]) (use n an in auto)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   179
    moreover assume "totient n \<noteq> n - 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   180
    then have "totient n > 0" "totient n < n - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   181
      using \<open>n \<ge> 2\<close> and totient_less[of n] by simp_all
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   182
    ultimately show False
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   183
      using nm by auto
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   184
  qed
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   185
qed (use n in auto)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
theorem lucas:
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   189
    and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> [a^((n - 1) div p) \<noteq> 1] (mod n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
  shows "prime n"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
proof-
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   192
  from n2 have n01: "n \<noteq> 0" "n \<noteq> 1" "n - 1 \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   193
    by arith+
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   194
  from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   195
    by simp
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 68157
diff changeset
   196
  from lucas_coprime_lemma[OF n01(3) an1] cong_imp_coprime an1
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   197
  have an: "coprime a n" "coprime (a ^ (n - 1)) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   198
    using \<open>n \<ge> 2\<close> by simp_all
67091
1393c2340eec more symbols;
wenzelm
parents: 67051
diff changeset
   199
  have False if H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "\<exists>m. ?P m")
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   200
  proof -
82690
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
   201
    from H0[unfolded exists_least_iff[of ?P]] obtain m where
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   202
      m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   203
      by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   204
    have False if nm1: "(n - 1) mod m > 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   205
    proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
      let ?y = "a^ ((n - 1) div m * m)"
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 63905
diff changeset
   208
      note mdeq = div_mult_mod_eq[of "(n - 1)" m]
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
      have yn: "coprime ?y n"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   210
        using an(1) by (cases "(n - Suc 0) div m * m = 0") auto
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
      have "?y mod n = (a^m)^((n - 1) div m) mod n"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
        by (simp add: algebra_simps power_mult)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
        using power_mod[of "a^m" n "(n - 1) div m"] by simp
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   215
      also have "\<dots> = 1" using m(3)[unfolded cong_def onen] onen
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
        by (metis power_one)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   217
      finally have *: "?y mod n = 1"  .
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   218
      have **: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   219
        using an1[unfolded cong_def onen] onen
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 63905
diff changeset
   220
          div_mult_mod_eq[of "(n - 1)" m, symmetric]
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   221
        by (simp add:power_add[symmetric] cong_def * del: One_nat_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   222
      have "[a ^ ((n - 1) mod m) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   223
        by (metis cong_mult_rcancel_nat mult.commute ** yn)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   224
      with m(4)[rule_format, OF th0] nm1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   225
        less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   226
        by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   227
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   228
    then have "(n - 1) mod m = 0" by auto
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
    then have mn: "m dvd n - 1" by presburger
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   230
    then obtain r where r: "n - 1 = m * r"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   231
      unfolding dvd_def by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   232
    from n01 r m(2) have r01: "r \<noteq> 0" "r \<noteq> 1" by auto
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
    obtain p where p: "prime p" "p dvd r"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
      by (metis prime_factor_nat r01(2))
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   235
    then have th: "prime p \<and> p dvd n - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   236
      unfolding r by (auto intro: dvd_mult)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   237
    from r have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
      by (simp add: power_mult)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   239
    also have "\<dots> = (a^(m*(r div p))) mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   240
      using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   241
    also have "\<dots> = ((a^m)^(r div p)) mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   242
      by (simp add: power_mult)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   243
    also have "\<dots> = ((a^m mod n)^(r div p)) mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   244
      using power_mod ..
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   245
    also from m(3) onen have "\<dots> = 1"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   246
      by (simp add: cong_def)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   248
      using onen by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   249
    with pn th show ?thesis by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   250
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   251
  then have "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   252
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   253
  then show ?thesis by (rule lucas_weak[OF n2 an1])
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   257
subsection \<open>Definition of the order of a number mod \<open>n\<close>\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   261
text \<open>This has the expected properties.\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
lemma coprime_ord:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   264
  fixes n::nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  assumes "coprime n a"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> [a^ m \<noteq> 1] (mod n))"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
proof-
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   269
  from bigger_prime[of a] obtain p where p: "prime p" "a < p"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   270
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   271
  from assms have o: "ord n a = Least ?P"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   272
    by (simp add: ord_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   273
  have ex: "\<exists>m>0. ?P m"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   274
  proof (cases "n \<ge> 2")
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   275
    case True
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   276
    moreover from assms have "coprime a n"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   277
      by (simp add: ac_simps)
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   278
    then have "[a ^ totient n = 1] (mod n)"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   279
      by (rule euler_theorem)
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   280
    ultimately show ?thesis
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   281
      by (auto intro: exI [where x = "totient n"])
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   282
  next
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   283
    case False
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   284
    then have "n = 0 \<or> n = 1"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   285
      by auto
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   286
    with assms show ?thesis
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   287
      by auto
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents: 65416
diff changeset
   288
  qed
82690
cccbfa567117 Sylvestre's correction to ex_least_nat_le and other tidying
paulson <lp15@cam.ac.uk>
parents: 82664
diff changeset
   289
  from exists_least_iff'[of ?P] ex assms show ?thesis
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    unfolding o[symmetric] by auto
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   293
text \<open>With the special value \<open>0\<close> for non-coprime case, it's more convenient.\<close>
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   294
lemma ord_works: "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   295
  for n :: nat
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   296
  by (cases "coprime n a") (use coprime_ord[of n a] in \<open>auto simp add: ord_def cong_def\<close>)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   298
lemma ord: "[a^(ord n a) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   299
  for n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   300
  using ord_works by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   302
lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> \<not> [a^m = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   303
  for n :: nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
  using ord_works by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   306
lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> \<not> coprime n a"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   307
  for n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   308
  by (cases "coprime n a") (simp add: coprime_ord, simp add: ord_def)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   310
lemma divides_rexp: "x dvd y \<Longrightarrow> x dvd (y ^ Suc n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   311
  for x y :: nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
  by (simp add: dvd_mult2[of x y])
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   314
lemma ord_divides:"[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   315
  (is "?lhs \<longleftrightarrow> ?rhs")
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   316
  for n :: nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
proof
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   318
  assume ?rhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   319
  then obtain k where "d = ord n a * k"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   320
    unfolding dvd_def by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   321
  then have "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   322
    by (simp add : cong_def power_mult power_mod)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   324
    using ord[of a n, unfolded cong_def]
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   325
    by (simp add: cong_def power_mod)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   326
  finally show ?lhs .
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
next
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   328
  assume ?lhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   329
  show ?rhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   330
  proof (cases "coprime n a")
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   331
    case prem: False
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   332
    then have o: "ord n a = 0" by (simp add: ord_def)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   333
    show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   334
    proof (cases d)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   335
      case 0
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   336
      with o prem show ?thesis by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   337
    next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   338
      case (Suc d')
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   339
      then have d0: "d \<noteq> 0" by simp
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   340
      from prem obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   341
        by (auto elim: not_coprimeE) 
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   342
      from \<open>?lhs\<close> obtain q1 q2 where q12: "a ^ d + n * q1 = 1 + n * q2"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   343
        using prem d0 lucas_coprime_lemma
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   344
        by (auto elim: not_coprimeE simp add: ac_simps)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   345
      then have "a ^ d + n * q1 - n * q2 = 1" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   346
      with dvd_diff_nat [OF dvd_add [OF divides_rexp]]  dvd_mult2 Suc p have "p dvd 1"
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   347
        by metis
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
      with p(3) have False by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   349
      then show ?thesis ..
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   350
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   351
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   352
    case H: True
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
    let ?o = "ord n a"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
    let ?q = "d div ord n a"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
    let ?r = "d mod ord n a"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
    have eqo: "[(a^?o)^?q = 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   357
      using cong_pow ord_works by fastforce
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
    from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
67345
debef21cbed6 tuned op's
nipkow
parents: 67091
diff changeset
   359
    then have opos: "?o > 0" by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   360
    from div_mult_mod_eq[of d "ord n a"] \<open>?lhs\<close>
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   361
    have "[a^(?o*?q + ?r) = 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   362
      by (simp add: cong_def mult.commute)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   363
    then have "[(a^?o)^?q * (a^?r) = 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   364
      by (simp add: cong_def power_mult[symmetric] power_add[symmetric])
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   365
    then have th: "[a^?r = 1] (mod n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
      using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   367
      by (simp add: cong_def del: One_nat_def) (metis mod_mult_left_eq nat_mult_1)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   368
    show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   369
    proof (cases "?r = 0")
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   370
      case True
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   371
      then show ?thesis by (simp add: dvd_eq_mod_eq_0)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   372
    next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   373
      case False
67345
debef21cbed6 tuned op's
nipkow
parents: 67091
diff changeset
   374
      with mod_less_divisor[OF opos, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   376
      show ?thesis by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   377
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   378
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   381
lemma order_divides_totient:
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   382
  "ord n a dvd totient n" if "coprime n a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   383
  using that euler_theorem [of a n]
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   384
  by (simp add: ord_divides [symmetric] ac_simps)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
lemma order_divides_expdiff:
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
  fixes n::nat and a::nat assumes na: "coprime n a"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
  shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   389
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   390
  have th: "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   391
    if na: "coprime n a" and ed: "(e::nat) \<le> d"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   392
    for n a d e :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   393
  proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   394
    from na ed have "\<exists>c. d = e + c" by presburger
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   395
    then obtain c where c: "d = e + c" ..
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
    from na have an: "coprime a n"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   397
      by (simp add: ac_simps)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   398
    then have aen: "coprime (a ^ e) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   399
      by (cases "e > 0") simp_all
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   400
    from an have acn: "coprime (a ^ c) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   401
      by (cases "c > 0") simp_all
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   402
    from c have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   403
      by simp
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
    also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
    also have  "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
      using cong_mult_lcancel_nat [OF aen, of "a^c" "a^0"] by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   407
    also have "\<dots> \<longleftrightarrow> ord n a dvd c"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   408
      by (simp only: ord_divides)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
    also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   410
      by (auto simp add: cong_altdef_nat)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   411
    finally show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   412
      using c by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   413
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   414
  consider "e \<le> d" | "d \<le> e" by arith
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   415
  then show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   416
  proof cases
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   417
    case 1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   418
    with na show ?thesis by (rule th)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   419
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   420
    case 2
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   421
    from th[OF na this] show ?thesis
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   422
      by (metis cong_sym)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   423
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   426
lemma ord_not_coprime [simp]: "\<not>coprime n a \<Longrightarrow> ord n a = 0"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   427
  by (simp add: ord_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   428
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   429
lemma ord_1 [simp]: "ord 1 n = 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   430
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   431
  have "(LEAST k. k > 0) = (1 :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   432
    by (rule Least_equality) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   433
  thus ?thesis by (simp add: ord_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   434
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   435
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   436
lemma ord_1_right [simp]: "ord (n::nat) 1 = 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   437
  using ord_divides[of 1 1 n] by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   438
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   439
lemma ord_Suc_0_right [simp]: "ord (n::nat) (Suc 0) = 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   440
  using ord_divides[of 1 1 n] by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   441
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   442
lemma ord_0_nat [simp]: "ord 0 (n :: nat) = (if n = 1 then 1 else 0)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   443
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   444
  have "(LEAST k. k > 0) = (1 :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   445
    by (rule Least_equality) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   446
  thus ?thesis by (auto simp: ord_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   447
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   448
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   449
lemma ord_0_right_nat [simp]: "ord (n :: nat) 0 = (if n = 1 then 1 else 0)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   450
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   451
  have "(LEAST k. k > 0) = (1 :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   452
    by (rule Least_equality) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   453
  thus ?thesis by (auto simp: ord_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   454
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   455
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   456
lemma ord_divides': "[a ^ d = Suc 0] (mod n) = (ord n a dvd d)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   457
  using ord_divides[of a d n] by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   458
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   459
lemma ord_Suc_0 [simp]: "ord (Suc 0) n = 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   460
  using ord_1[where 'a = nat] by (simp del: ord_1)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   461
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   462
lemma ord_mod [simp]: "ord n (k mod n) = ord n k"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   463
  by (cases "n = 0") (auto simp add: ord_def cong_def power_mod)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   464
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   465
lemma ord_gt_0_iff [simp]: "ord (n::nat) x > 0 \<longleftrightarrow> coprime n x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   466
  using ord_eq_0[of n x] by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   467
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   468
lemma ord_eq_Suc_0_iff: "ord n (x::nat) = Suc 0 \<longleftrightarrow> [x = 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   469
  using ord_divides[of x 1 n] by (auto simp: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   470
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   471
lemma ord_cong:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   472
  assumes "[k1 = k2] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   473
  shows   "ord n k1 = ord n k2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   474
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   475
  have "ord n (k1 mod n) = ord n (k2 mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   476
    by (simp only: assms[unfolded cong_def])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   477
  thus ?thesis by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   478
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   479
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   480
lemma ord_nat_code [code_unfold]:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   481
  "ord n a =
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   482
     (if n = 0 then if a = 1 then 1 else 0 else
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   483
        if coprime n a then Min (Set.filter (\<lambda>k. [a ^ k = 1] (mod n)) {0<..n}) else 0)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   484
proof (cases "coprime n a \<and> n > 0")
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   485
  case True
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   486
  define A where "A = {k\<in>{0<..n}. [a ^ k = 1] (mod n)}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   487
  define k where "k = (LEAST k. k > 0 \<and> [a ^ k = 1] (mod n))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   488
  have totient: "totient n \<in> A"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   489
    using euler_theorem[of a n] True
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   490
    by (auto simp: A_def coprime_commute intro!: Nat.gr0I totient_le)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   491
  moreover have "finite A" by (auto simp: A_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   492
  ultimately have *: "Min A \<in> A" and "\<forall>y. y \<in> A \<longrightarrow> Min A \<le> y"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   493
    by (auto intro: Min_in)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   494
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   495
  have "k > 0 \<and> [a ^ k = 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   496
    unfolding k_def by (rule LeastI[of _ "totient n"]) (use totient in \<open>auto simp: A_def\<close>)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   497
  moreover have "k \<le> totient n"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   498
    unfolding k_def by (intro Least_le) (use totient in \<open>auto simp: A_def\<close>)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   499
  ultimately have "k \<in> A" using totient_le[of n] by (auto simp: A_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   500
  hence "Min A \<le> k" by (intro Min_le) (auto simp: \<open>finite A\<close>)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   501
  moreover from * have "k \<le> Min A"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   502
    unfolding k_def by (intro Least_le) (auto simp: A_def)
82664
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 78668
diff changeset
   503
  ultimately show ?thesis using True
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 78668
diff changeset
   504
    by (simp add: ord_def k_def A_def)
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   505
qed auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   506
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   507
theorem ord_modulus_mult_coprime:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   508
  fixes x :: nat
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   509
  assumes "coprime m n"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   510
  shows   "ord (m * n) x = lcm (ord m x) (ord n x)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   511
proof (intro dvd_antisym)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   512
  have "[x ^ lcm (ord m x) (ord n x) = 1] (mod (m * n))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   513
    using assms by (intro coprime_cong_mult_nat assms) (auto simp: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   514
  thus "ord (m * n) x dvd lcm (ord m x) (ord n x)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   515
    by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   516
next
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   517
  show "lcm (ord m x) (ord n x) dvd ord (m * n) x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   518
  proof (intro lcm_least)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   519
    show "ord m x dvd ord (m * n) x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   520
      using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 m n] assms
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   521
      by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   522
    show "ord n x dvd ord (m * n) x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   523
      using cong_modulus_mult_nat[of "x ^ ord (m * n) x" 1 n m] assms
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   524
      by (simp add: ord_divides' mult.commute)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   525
  qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   526
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   527
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   528
corollary ord_modulus_prod_coprime:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   529
  assumes "finite A" "\<And>i j. i \<in> A \<Longrightarrow> j \<in> A \<Longrightarrow> i \<noteq> j \<Longrightarrow> coprime (f i) (f j)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   530
  shows   "ord (\<Prod>i\<in>A. f i :: nat) x = (LCM i\<in>A. ord (f i) x)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   531
  using assms by (induction A rule: finite_induct)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   532
                 (simp, simp, subst ord_modulus_mult_coprime, auto intro!: prod_coprime_right)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   533
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   534
lemma ord_power_aux:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   535
  fixes m x k a :: nat
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   536
  defines "l \<equiv> ord m a"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   537
  shows   "ord m (a ^ k) * gcd k l = l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   538
proof (rule dvd_antisym)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   539
  have "[a ^ lcm k l = 1] (mod m)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   540
    unfolding ord_divides by (simp add: l_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   541
  also have "lcm k l = k * (l div gcd k l)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   542
    by (simp add: lcm_nat_def div_mult_swap)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   543
  finally have "ord m (a ^ k) dvd l div gcd k l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   544
    unfolding ord_divides [symmetric] by (simp add: power_mult [symmetric])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   545
  thus "ord m (a ^ k) * gcd k l dvd l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   546
    by (cases "l = 0") (auto simp: dvd_div_iff_mult)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   547
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   548
  have "[(a ^ k) ^ ord m (a ^ k) = 1] (mod m)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   549
    by (rule ord)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   550
  also have "(a ^ k) ^ ord m (a ^ k) = a ^ (k * ord m (a ^ k))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   551
    by (simp add: power_mult)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   552
  finally have "ord m a dvd k * ord m (a ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   553
    by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   554
  hence "l dvd gcd (k * ord m (a ^ k)) (l * ord m (a ^ k))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   555
    by (intro gcd_greatest dvd_triv_left) (auto simp: l_def ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   556
  also have "gcd (k * ord m (a ^ k)) (l * ord m (a ^ k)) = ord m (a ^ k) * gcd k l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   557
    by (subst gcd_mult_distrib_nat) (auto simp: mult_ac)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   558
  finally show "l dvd ord m (a ^ k) * gcd k l" .
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   559
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   560
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   561
theorem ord_power: "coprime m a \<Longrightarrow> ord m (a ^ k :: nat) = ord m a div gcd k (ord m a)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   562
  using ord_power_aux[of m a k] by (metis div_mult_self_is_m gcd_pos_nat ord_eq_0)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   563
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   564
lemma inj_power_mod:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   565
  assumes "coprime n (a :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   566
  shows   "inj_on (\<lambda>k. a ^ k mod n) {..<ord n a}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   567
proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   568
  fix k l assume *: "k \<in> {..<ord n a}" "l \<in> {..<ord n a}" "a ^ k mod n = a ^ l mod n"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   569
  have "k = l" if "k < l" "l < ord n a" "[a ^ k = a ^ l] (mod n)" for k l
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   570
  proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   571
    have "l = k + (l - k)" using that by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   572
    also have "a ^ \<dots> = a ^ k * a ^ (l - k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   573
      by (simp add: power_add)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   574
    also have "[\<dots> = a ^ l * a ^ (l - k)] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   575
      using that by (intro cong_mult) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   576
    finally have "[a ^ l * a ^ (l - k) = a ^ l * 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   577
      by (simp add: cong_sym_eq)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   578
    with assms have "[a ^ (l - k) = 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   579
      by (subst (asm) cong_mult_lcancel_nat) (auto simp: coprime_commute)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   580
    hence "ord n a dvd l - k"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   581
      by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   582
    from dvd_imp_le[OF this] and \<open>l < ord n a\<close> have "l - k = 0"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   583
      by (cases "l - k = 0") auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   584
    with \<open>k < l\<close> show "k = l" by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   585
  qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   586
  from this[of k l] and this[of l k] and * show "k = l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   587
    by (cases k l rule: linorder_cases) (auto simp: cong_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   588
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   589
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   590
lemma ord_eq_2_iff: "ord n (x :: nat) = 2 \<longleftrightarrow> [x \<noteq> 1] (mod n) \<and> [x\<^sup>2 = 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   591
proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   592
  assume x: "[x \<noteq> 1] (mod n) \<and> [x\<^sup>2 = 1] (mod n)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   593
  hence "coprime n x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   594
    by (metis coprime_commute lucas_coprime_lemma zero_neq_numeral)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   595
  with x have "ord n x dvd 2" "ord n x \<noteq> 1" "ord n x > 0"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   596
    by (auto simp: ord_divides' ord_eq_Suc_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   597
  thus "ord n x = 2" by (auto dest!: dvd_imp_le simp del: ord_gt_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   598
qed (use ord_divides[of _ 2] ord_divides[of _ 1] in auto)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   599
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   600
lemma square_mod_8_eq_1_iff: "[x\<^sup>2 = 1] (mod 8) \<longleftrightarrow> odd (x :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   601
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   602
  have "[x\<^sup>2 = 1] (mod 8) \<longleftrightarrow> ((x mod 8)\<^sup>2 mod 8 = 1)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   603
    by (simp add: power_mod cong_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   604
  also have "\<dots> \<longleftrightarrow> x mod 8 \<in> {1, 3, 5, 7}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   605
  proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   606
    assume x: "(x mod 8)\<^sup>2 mod 8 = 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   607
    have "x mod 8 \<in> {..<8}" by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   608
    also have "{..<8} = {0, 1, 2, 3, 4, 5, 6, 7::nat}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   609
      by (simp add: lessThan_nat_numeral lessThan_Suc insert_commute)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   610
    finally have x_cases: "x mod 8 \<in> {0, 1, 2, 3, 4, 5, 6, 7}" .
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   611
    from x have "x mod 8 \<notin> {0, 2, 4, 6}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   612
      using x by (auto intro: Nat.gr0I)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   613
    with x_cases show "x mod 8 \<in> {1, 3, 5, 7}" by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   614
  qed auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   615
  also have "\<dots> \<longleftrightarrow> odd (x mod 8)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   616
    by (auto elim!: oddE)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   617
  also have "\<dots> \<longleftrightarrow> odd x"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   618
    by presburger
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   619
  finally show ?thesis .
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   620
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   621
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   622
lemma ord_twopow_aux:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   623
  assumes "k \<ge> 3" and "odd (x :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   624
  shows   "[x ^ (2 ^ (k - 2)) = 1] (mod (2 ^ k))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   625
  using assms(1)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   626
proof (induction k rule: dec_induct)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   627
  case base
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   628
  from assms have "[x\<^sup>2 = 1] (mod 8)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   629
    by (subst square_mod_8_eq_1_iff) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   630
  thus ?case by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   631
next
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   632
  case (step k)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   633
  define k' where "k' = k - 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   634
  have k: "k = Suc (Suc k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   635
    using \<open>k \<ge> 3\<close> by (simp add: k'_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   636
  from \<open>k \<ge> 3\<close> have "2 * k \<ge> Suc k" by presburger
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   637
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   638
  from \<open>odd x\<close> have "x > 0" by (intro Nat.gr0I) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   639
  from step.IH have "2 ^ k dvd (x ^ (2 ^ (k - 2)) - 1)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   640
    by (rule cong_to_1_nat)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   641
  then obtain t where "x ^ (2 ^ (k - 2)) - 1 = t * 2 ^ k"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   642
    by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   643
  hence "x ^ (2 ^ (k - 2)) = t * 2 ^ k + 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   644
    by (metis \<open>0 < x\<close> add.commute add_diff_inverse_nat less_one neq0_conv power_eq_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   645
  hence "(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   646
    by (rule arg_cong)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   647
  hence "[(x ^ (2 ^ (k - 2))) ^ 2 = (t * 2 ^ k + 1) ^ 2] (mod (2 ^ Suc k))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   648
    by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   649
  also have "(x ^ (2 ^ (k - 2))) ^ 2 = x ^ (2 ^ (k - 1))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   650
    by (simp_all add: power_even_eq[symmetric] power_mult k )
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   651
  also have "(t * 2 ^ k + 1) ^ 2 = t\<^sup>2 * 2 ^ (2 * k) + t * 2 ^ Suc k + 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   652
    by (subst power2_eq_square)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   653
       (auto simp: algebra_simps k power2_eq_square[of t]
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   654
                   power_even_eq[symmetric] power_add [symmetric])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   655
  also have "[\<dots> = 0 + 0 + 1] (mod 2 ^ Suc k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   656
    using \<open>2 * k \<ge> Suc k\<close>
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   657
    by (intro cong_add)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   658
       (auto simp: cong_0_iff intro: dvd_mult[OF le_imp_power_dvd] simp del: power_Suc)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   659
  finally show ?case by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   660
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   661
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   662
lemma ord_twopow_3_5:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   663
  assumes "k \<ge> 3" "x mod 8 \<in> {3, 5 :: nat}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   664
  shows   "ord (2 ^ k) x = 2 ^ (k - 2)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   665
  using assms(1)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   666
proof (induction k rule: less_induct)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   667
  have "x mod 8 = 3 \<or> x mod 8 = 5" using assms by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   668
  hence "odd x" by presburger
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   669
  case (less k)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   670
  from \<open>k \<ge> 3\<close> consider "k = 3" | "k = 4" | "k \<ge> 5" by force
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   671
  thus ?case
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   672
  proof cases
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   673
    case 1
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   674
    thus ?thesis using assms
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   675
      by (auto simp: ord_eq_2_iff cong_def simp flip: power_mod[of x])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   676
  next
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   677
    case 2
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   678
    from assms have "x mod 8 = 3 \<or> x mod 8 = 5" by auto
76231
8a48e18f081e reduce prominence of facts
haftmann
parents: 69785
diff changeset
   679
    then have x': "x mod 16 = 3 \<or> x mod 16 = 5 \<or> x mod 16 = 11 \<or> x mod 16 = 13"
78668
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 76231
diff changeset
   680
      using mod_double_nat [of x 8] by auto
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   681
    hence "[x ^ 4 = 1] (mod 16)" using assms
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   682
      by (auto simp: cong_def simp flip: power_mod[of x])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   683
    hence "ord 16 x dvd 2\<^sup>2" by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   684
    then obtain l where l: "ord 16 x = 2 ^ l" "l \<le> 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   685
      by (subst (asm) divides_primepow_nat) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   686
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   687
    have "[x ^ 2 \<noteq> 1] (mod 16)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   688
      using x' by (auto simp: cong_def simp flip: power_mod[of x])
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   689
    hence "\<not>ord 16 x dvd 2" by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   690
    with l have "l = 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   691
      using le_imp_power_dvd[of l 1 2] by (cases "l \<le> 1") auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   692
    with l show ?thesis by (simp add: \<open>k = 4\<close>)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   693
  next
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   694
    case 3
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   695
    define k' where "k' = k - 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   696
    have k': "k' \<ge> 2" and [simp]: "k = Suc (Suc k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   697
      using 3 by (simp_all add: k'_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   698
    have IH: "ord (2 ^ k') x = 2 ^ (k' - 2)" "ord (2 ^ Suc k') x = 2 ^ (k' - 1)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   699
      using less.IH[of k'] less.IH[of "Suc k'"] 3 by simp_all
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   700
    from IH have cong: "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ k'))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   701
      by (simp_all add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   702
    have notcong: "[x ^ (2 ^ (k' - 2)) \<noteq> 1] (mod (2 ^ Suc k'))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   703
    proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   704
      assume "[x ^ (2 ^ (k' - 2)) = 1] (mod (2 ^ Suc k'))"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   705
      hence "ord (2 ^ Suc k') x dvd 2 ^ (k' - 2)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   706
        by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   707
      also have "ord (2 ^ Suc k') x = 2 ^ (k' - 1)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   708
        using IH by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   709
      finally have "k' - 1 \<le> k' - 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   710
        by (rule power_dvd_imp_le) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   711
      with \<open>k' \<ge> 2\<close> show False by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   712
    qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   713
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   714
    have "2 ^ k' + 1 < 2 ^ k' + (2 ^ k' :: nat)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   715
      using one_less_power[of "2::nat" k'] k' by (intro add_strict_left_mono) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   716
    with cong notcong have cong': "x ^ (2 ^ (k' - 2)) mod 2 ^ Suc k' = 1 + 2 ^ k'"
78668
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 76231
diff changeset
   717
      using mod_double_nat [of \<open>x ^ 2 ^ (k' - 2)\<close> \<open>2 ^ k'\<close>] k' by (auto simp: cong_def)
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   718
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   719
    hence "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' \<or>
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   720
           x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'"
78668
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 76231
diff changeset
   721
      using mod_double_nat [of \<open>x ^ 2 ^ (k' - 2)\<close> \<open>2 ^ Suc k'\<close>] by auto
69785
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   722
    hence eq: "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   723
    proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   724
      assume *: "x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   725
      have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   726
        by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   727
      also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 2 ^ k'] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   728
        by (subst *) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   729
      finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   730
        by (rule cong_pow)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   731
      hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 2 ^ k') ^ 2] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   732
        by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   733
      also have "Suc (k' - 2) = k' - 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   734
        using k' by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   735
      also have "(1 + 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ (2 * k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   736
        by (subst power2_eq_square) (simp add: algebra_simps flip: power_add)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   737
      also have "(2 ^ k :: nat) dvd 2 ^ (2 * k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   738
        using k' by (intro le_imp_power_dvd) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   739
      hence "[1 + 2 ^ (k - 1) + 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + (0 :: nat)] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   740
        by (intro cong_add) (auto simp: cong_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   741
      finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   742
        by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   743
    next
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   744
      assume *: "x ^ (2 ^ (k' - 2)) mod 2 ^ k = 1 + 2 ^ k' + 2 ^ Suc k'"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   745
      have "[x ^ (2 ^ (k' - 2)) = x ^ (2 ^ (k' - 2)) mod 2 ^ k] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   746
        by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   747
      also have "[x ^ (2 ^ (k' - 2)) mod (2 ^ k) = 1 + 3 * 2 ^ k'] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   748
        by (subst *) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   749
      finally have "[(x ^ 2 ^ (k' - 2)) ^ 2 = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   750
        by (rule cong_pow)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   751
      hence "[x ^ 2 ^ Suc (k' - 2) = (1 + 3 * 2 ^ k') ^ 2] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   752
        by (simp add: power_mult [symmetric] power_Suc2 [symmetric] del: power_Suc)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   753
      also have "Suc (k' - 2) = k' - 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   754
        using k' by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   755
      also have "(1 + 3 * 2 ^ k' :: nat)\<^sup>2 = 1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   756
        by (subst power2_eq_square) (simp add: algebra_simps flip: power_add)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   757
      also have "(2 ^ k :: nat) dvd 9 * 2 ^ (2 * k')"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   758
        using k' by (intro dvd_mult le_imp_power_dvd) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   759
      hence "[1 + 2 ^ (k - 1) + 2 ^ k + 9 * 2 ^ (2 * k') = 1 + 2 ^ (k - 1) + 0 + (0 :: nat)]
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   760
               (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   761
        by (intro cong_add) (auto simp: cong_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   762
      finally show "[x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1)] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   763
        by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   764
    qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   765
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   766
    have notcong': "[x ^ 2 ^ (k - 3) \<noteq> 1] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   767
    proof
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   768
      assume "[x ^ 2 ^ (k - 3) = 1] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   769
      hence "[x ^ 2 ^ (k' - 1) - x ^ 2 ^ (k' - 1) = 1 + 2 ^ (k - 1) - 1] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   770
        by (intro cong_diff_nat eq) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   771
      hence "[2 ^ (k - 1) = (0 :: nat)] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   772
        by (simp add: cong_sym_eq)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   773
      hence "2 ^ k dvd 2 ^ (k - 1)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   774
        by (simp add: cong_0_iff)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   775
      hence "k \<le> k - 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   776
        by (rule power_dvd_imp_le) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   777
      thus False by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   778
    qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   779
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   780
    have "[x ^ 2 ^ (k - 2) = 1] (mod 2 ^ k)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   781
      using ord_twopow_aux[of k x] \<open>odd x\<close> \<open>k \<ge> 3\<close> by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   782
    hence "ord (2 ^ k) x dvd 2 ^ (k - 2)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   783
      by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   784
    then obtain l where l: "l \<le> k - 2" "ord (2 ^ k) x = 2 ^ l"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   785
      using divides_primepow_nat[of 2 "ord (2 ^ k) x" "k - 2"] by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   786
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   787
    from notcong' have "\<not>ord (2 ^ k) x dvd 2 ^ (k - 3)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   788
      by (simp add: ord_divides')
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   789
    with l have "l = k - 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   790
      using le_imp_power_dvd[of l "k - 3" 2] by (cases "l \<le> k - 3") auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   791
    with l show ?thesis by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   792
  qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   793
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   794
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   795
lemma ord_4_3 [simp]: "ord 4 (3::nat) = 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   796
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   797
  have "[3 ^ 2 = (1 :: nat)] (mod 4)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   798
    by (simp add: cong_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   799
  hence "ord 4 (3::nat) dvd 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   800
    by (subst (asm) ord_divides) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   801
  hence "ord 4 (3::nat) \<le> 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   802
    by (intro dvd_imp_le) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   803
  moreover have "ord 4 (3::nat) \<noteq> 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   804
    by (auto simp: ord_eq_Suc_0_iff cong_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   805
  moreover have "ord 4 (3::nat) \<noteq> 0"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   806
    by (auto simp: gcd_non_0_nat coprime_iff_gcd_eq_1)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   807
  ultimately show "ord 4 (3 :: nat) = 2"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   808
    by linarith
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   809
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   810
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   811
lemma elements_with_ord_1: "n > 0 \<Longrightarrow> {x\<in>totatives n. ord n x = Suc 0} = {1}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   812
  by (auto simp: ord_eq_Suc_0_iff cong_def totatives_less)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   813
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   814
lemma residue_prime_has_primroot:
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   815
  fixes p :: nat
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   816
  assumes "prime p"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   817
  shows "\<exists>a\<in>totatives p. ord p a = p - 1"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   818
proof -
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   819
  from residue_prime_mult_group_has_gen[OF assms]
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   820
    obtain a where a: "a \<in> {1..p-1}" "{1..p-1} = {a ^ i mod p |i. i \<in> UNIV}" by blast
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   821
  from a have "coprime p a"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   822
    using a assms by (intro prime_imp_coprime) (auto dest: dvd_imp_le)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   823
  with a(1) have "a \<in> totatives p" by (auto simp: totatives_def coprime_commute)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   824
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   825
  have "p - 1 = card {1..p-1}" by simp
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   826
  also have "{1..p-1} = {a ^ i mod p |i. i \<in> UNIV}" by fact
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   827
  also have "{a ^ i mod p |i. i \<in> UNIV} = (\<lambda>i. a ^ i mod p) ` {..<ord p a}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   828
  proof (intro equalityI subsetI)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   829
    fix x assume "x \<in> {a ^ i mod p |i. i \<in> UNIV}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   830
    then obtain i where [simp]: "x = a ^ i mod p" by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   831
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   832
    have "[a ^ i = a ^ (i mod ord p a)] (mod p)"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   833
      using \<open>coprime p a\<close> by (subst order_divides_expdiff) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   834
    hence "\<exists>j. a ^ i mod p = a ^ j mod p \<and> j < ord p a"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   835
      using \<open>coprime p a\<close> by (intro exI[of _ "i mod ord p a"]) (auto simp: cong_def)
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   836
    thus "x \<in> (\<lambda>i. a ^ i mod p) ` {..<ord p a}"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   837
      by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   838
  qed auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   839
  also have "card \<dots> = ord p a"
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   840
    using inj_power_mod[OF \<open>coprime p a\<close>] by (subst card_image) auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   841
  finally show ?thesis using \<open>a \<in> totatives p\<close>
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   842
    by auto
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   843
qed
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   844
9e326f6f8a24 More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   845
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   846
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   847
subsection \<open>Another trivial primality characterization\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   849
lemma prime_prime_factor: "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   850
  (is "?lhs \<longleftrightarrow> ?rhs")
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   851
  for n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   852
proof (cases "n = 0 \<or> n = 1")
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   853
  case True
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   854
  then show ?thesis
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
   855
     by (metis bigger_prime dvd_0_right not_prime_1 not_prime_0)
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   856
next
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   857
  case False
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   858
  show ?thesis
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   859
  proof
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   860
    assume "prime n"
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   861
    then show ?rhs
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   862
      by (metis not_prime_1 prime_nat_iff)
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   863
  next
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   864
    assume ?rhs
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   865
    with False show "prime n"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
   866
      by (auto simp: prime_nat_iff) (metis One_nat_def prime_factor_nat prime_nat_iff)
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   867
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   870
lemma prime_divisor_sqrt: "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d\<^sup>2 \<le> n \<longrightarrow> d = 1)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   871
  for n :: nat
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
proof -
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   873
  consider "n = 0" | "n = 1" | "n \<noteq> 0" "n \<noteq> 1" by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   874
  then show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   875
  proof cases
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   876
    case 1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   877
    then show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   878
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   879
    case 2
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   880
    then show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   881
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   882
    case n: 3
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   883
    then have np: "n > 1" by arith
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   884
    {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   885
      fix d
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   886
      assume d: "d dvd n" "d\<^sup>2 \<le> n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   887
        and H: "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   888
      from H d have d1n: "d = 1 \<or> d = n" by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   889
      then have "d = 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   890
      proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   891
        assume dn: "d = n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   892
        from n have "n\<^sup>2 > n * 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   893
          by (simp add: power2_eq_square)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   894
        with dn d(2) show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   895
      qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   896
    }
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
    moreover
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   898
    {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   899
      fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'\<^sup>2 \<le> n \<longrightarrow> d' = 1"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
      from d n have "d \<noteq> 0"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
        by (metis dvd_0_left_iff)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   902
      then have dp: "d > 0" by simp
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
      from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
      from n dp e have ep:"e > 0" by simp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   905
      from dp ep have "d\<^sup>2 \<le> n \<or> e\<^sup>2 \<le> n"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
        by (auto simp add: e power2_eq_square mult_le_cancel_left)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   907
      then have "d = 1 \<or> d = n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   908
      proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   909
        assume "d\<^sup>2 \<le> n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   910
        with H[rule_format, of d] d have "d = 1" by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   911
        then show ?thesis ..
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   912
      next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   913
        assume h: "e\<^sup>2 \<le> n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   914
        from e have "e dvd n" by (simp add: dvd_def mult.commute)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   915
        with H[rule_format, of e] h have "e = 1" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   916
        with e have "d = n" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   917
        then show ?thesis ..
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   918
      qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   919
    }
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   920
    ultimately show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   921
      unfolding prime_nat_iff using np n(2) by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   922
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
lemma prime_prime_factor_sqrt:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   926
  "prime (n::nat) \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> (\<nexists>p. prime p \<and> p dvd n \<and> p\<^sup>2 \<le> n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
  (is "?lhs \<longleftrightarrow>?rhs")
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   928
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   929
  consider "n = 0" | "n = 1" | "n \<noteq> 0" "n \<noteq> 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   930
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   931
  then show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   932
  proof cases
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   933
    case 1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   934
    then show ?thesis by (metis not_prime_0)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   935
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   936
    case 2
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   937
    then show ?thesis by (metis not_prime_1)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   938
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   939
    case n: 3
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   940
    show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   941
    proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   942
      assume ?lhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   943
      from this[unfolded prime_divisor_sqrt] n show ?rhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   944
        by (metis prime_prime_factor)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   945
    next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   946
      assume ?rhs
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   947
      {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   948
        fix d
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   949
        assume d: "d dvd n" "d\<^sup>2 \<le> n" "d \<noteq> 1"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
        then obtain p where p: "prime p" "p dvd d"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   951
          by (metis prime_factor_nat)
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
   952
        from d(1) n have dp: "d > 0"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   953
          by (metis dvd_0_left neq0_conv)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
        from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
        have "p\<^sup>2 \<le> n" unfolding power2_eq_square by arith
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   956
        with \<open>?rhs\<close> n p(1) dvd_trans[OF p(2) d(1)] have False
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   957
          by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   958
      }
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   959
      with n prime_divisor_sqrt show ?lhs by auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   960
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   961
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   965
subsection \<open>Pocklington theorem\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
lemma pocklington_lemma:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   968
  fixes p :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   969
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   970
    and an: "[a^ (n - 1) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   971
    and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   972
    and pp: "prime p" and pn: "p dvd n"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
  shows "[p = 1] (mod q)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
proof -
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   975
  have p01: "p \<noteq> 0" "p \<noteq> 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   976
    using pp by (auto intro: prime_gt_0_nat)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   977
  obtain k where k: "a ^ (q * r) - 1 = n * k"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
    by (metis an cong_to_1_nat dvd_def nqr)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   979
  from pn[unfolded dvd_def] obtain l where l: "n = p * l"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   980
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   981
  have a0: "a \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   982
  proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   983
    assume "a = 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   984
    with n have "a^ (n - 1) = 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   985
      by (simp add: power_0_left)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   986
    with n an mod_less[of 1 n] show False
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
   987
      by (simp add: power_0_left cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   988
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   989
  with n nqr have aqr0: "a ^ (q * r) \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   990
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   991
  then have "(a ^ (q * r) - 1) + 1  = a ^ (q * r)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   992
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   993
  with k l have "a ^ (q * r) = p * l * k + 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   994
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   995
  then have "a ^ (r * q) + p * 0 = 1 + p * (l * k)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   996
    by (simp add: ac_simps)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   997
  then have odq: "ord p (a^r) dvd q"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
    unfolding ord_divides[symmetric] power_mult[symmetric]
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   999
    by (metis an cong_dvd_modulus_nat mult.commute nqr pn)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1000
  from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1001
    by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1002
  have d1: "d = 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1003
  proof (rule ccontr)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1004
    assume d1: "d \<noteq> 1"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
    obtain P where P: "prime P" "P dvd d"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1006
      by (metis d1 prime_factor_nat)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
    from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
    from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
    from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1010
    from P(1) have P0: "P \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1011
      by (metis not_prime_0)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
    from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
    from d s t P0  have s': "ord p (a^r) * t = s"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1014
      by (metis mult.commute mult_cancel1 mult.assoc)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
    have "ord p (a^r) * t*r = r * ord p (a^r) * t"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55370
diff changeset
  1016
      by (metis mult.assoc mult.commute)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1017
    then have exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
      by (simp only: power_mult)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1019
    then have "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1020
      by (metis cong_pow ord power_one)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1021
    then have pd0: "p dvd a^(ord p (a^r) * t*r) - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1022
      by (metis cong_to_1_nat exps)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1023
    from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1024
      using P0 by simp
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1025
    with caP have "coprime (a ^ (ord p (a ^ r) * t * r) - 1) n"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1026
      by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1027
    with p01 pn pd0 coprime_common_divisor [of _ n p] show False
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1028
      by auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1029
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1030
  with d have o: "ord p (a^r) = q" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1031
  from pp totient_prime [of p] have totient_eq: "totient p = p - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1032
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1033
  {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1034
    fix d
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1035
    assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63534
diff changeset
  1036
    from pp[unfolded prime_nat_iff] d have dp: "d = p" by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
    from n have "n \<noteq> 0" by simp
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1038
    then have False using d dp pn an
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1039
      by auto (metis One_nat_def Suc_lessI
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1040
        \<open>1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)\<close> \<open>a ^ (q * r) = p * l * k + 1\<close> add_diff_cancel_left' dvd_diff_nat dvd_power dvd_triv_left gcd_nat.trans nat_dvd_not_less nqr zero_less_diff zero_less_one) 
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1041
  }
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1042
  then have cpa: "coprime p a"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1043
    by (auto intro: coprimeI)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1044
  then have arp: "coprime (a ^ r) p"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
  1045
    by (cases "r > 0") (simp_all add: ac_simps)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1046
  from euler_theorem [OF arp, simplified ord_divides] o totient_eq have "q dvd (p - 1)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1047
    by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1048
  then obtain d where d:"p - 1 = q * d"
55337
5d45fb978d5a Number_Theory no longer introduces One_nat_def as a simprule. Tidied some proofs.
paulson <lp15@cam.ac.uk>
parents: 55321
diff changeset
  1049
    unfolding dvd_def by blast
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1050
  have "p \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1051
    by (metis p01(1))
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1052
  with d have "p + q * 0 = 1 + q * d" by simp
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
  then show ?thesis
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55370
diff changeset
  1054
    by (metis cong_iff_lin_nat mult.commute)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
theorem pocklington:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1058
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r" and sqr: "n \<le> q\<^sup>2"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1059
    and an: "[a^ (n - 1) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1060
    and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
  shows "prime n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1062
  unfolding prime_prime_factor_sqrt[of n]
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1063
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1064
  let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> (\<nexists>p. prime p \<and> p dvd n \<and> p\<^sup>2 \<le> n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1065
  from n have n01: "n \<noteq> 0" "n \<noteq> 1" by arith+
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1066
  {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1067
    fix p
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1068
    assume p: "prime p" "p dvd n" "p\<^sup>2 \<le> n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1069
    from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1070
      by (simp add: power2_eq_square)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1071
    then have pq: "p \<le> q"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1072
      by (metis le0 power_le_imp_le_base)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1073
    from pocklington_lemma[OF n nqr an aq p(1,2)] have *: "q dvd p - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1074
      by (metis cong_to_1_nat)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1075
    have "p - 1 \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1076
      using prime_ge_2_nat [OF p(1)] by arith
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1077
    with pq * have False
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1078
      by (simp add: nat_dvd_not_less)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1079
  }
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
  with n01 show ?ths by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1083
text \<open>Variant for application, to separate the exponentiation.\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
lemma pocklington_alt:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1085
  assumes n: "n \<ge> 2" and nqr: "n - 1 = q * r" and sqr: "n \<le> q\<^sup>2"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1086
    and an: "[a^ (n - 1) = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1087
    and aq: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
  shows "prime n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1089
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1090
  {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1091
    fix p
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1092
    assume p: "prime p" "p dvd q"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1093
    from aq[rule_format] p obtain b where b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1094
      by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1095
    have a0: "a \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1096
    proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1097
      assume a0: "a = 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1098
      from n an have "[0 = 1] (mod n)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1099
        unfolding a0 power_0_left by auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1100
      then show False
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1101
        using n by (simp add: cong_def dvd_eq_mod_eq_0[symmetric])
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1102
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1103
    then have a1: "a \<ge> 1" by arith
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
    from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1105
    have b0: "b \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1106
    proof
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1107
      assume b0: "b = 0"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
      from p(2) nqr have "(n - 1) mod p = 0"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
        by (metis mod_0 mod_mod_cancel mod_mult_self1_is_0)
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 63905
diff changeset
  1110
      with div_mult_mod_eq[of "n - 1" p]
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
      have "(n - 1) div p * p= n - 1" by auto
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1112
      then have eq: "(a^((n - 1) div p))^p = a^(n - 1)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
        by (simp only: power_mult[symmetric])
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1114
      have "p - 1 \<noteq> 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1115
        using prime_ge_2_nat [OF p(1)] by arith
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
      then have pS: "Suc (p - 1) = p" by arith
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1117
      from b have d: "n dvd a^((n - 1) div p)"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1118
        unfolding b0 by auto
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1119
      from divides_rexp[OF d, of "p - 1"] pS eq cong_dvd_iff [OF an] n show False
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1120
        by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1121
    qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1122
    then have b1: "b \<ge> 1" by arith
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 68157
diff changeset
  1123
    from cong_imp_coprime[OF Cong.cong_diff_nat[OF cong_sym [OF b(1)] cong_refl [of 1] b1]]
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1124
      ath b1 b nqr
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
    have "coprime (a ^ ((n - 1) div p) - 1) n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1126
      by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1127
  }
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1128
  then have "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
    by blast
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1130
  then show ?thesis by (rule pocklington[OF n nqr sqr an])
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1134
subsection \<open>Prime factorizations\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
55370
e6be866b5f5b minimal document;
wenzelm
parents: 55346
diff changeset
  1136
(* FIXME some overlap with material in UniqueFactorization, class unique_factorization *)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1138
definition "primefact ps n \<longleftrightarrow> foldr (*) ps 1 = n \<and> (\<forall>p\<in> set ps. prime p)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1140
lemma primefact:
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1141
  fixes n :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1142
  assumes n: "n \<noteq> 0"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
  shows "\<exists>ps. primefact ps n"
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62429
diff changeset
  1144
proof -
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1145
  obtain xs where xs: "mset xs = prime_factorization n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1146
    using ex_mset [of "prime_factorization n"] by blast
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1147
  from assms have "n = prod_mset (prime_factorization n)"
63830
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63633
diff changeset
  1148
    by (simp add: prod_mset_prime_factorization)
2ea3725a34bd msetsum -> set_mset, msetprod -> prod_mset
nipkow
parents: 63633
diff changeset
  1149
  also have "\<dots> = prod_mset (mset xs)" by (simp add: xs)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1150
  also have "\<dots> = foldr (*) xs 1" by (induct xs) simp_all
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1151
  finally have "foldr (*) xs 1 = n" ..
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1152
  moreover from xs have "\<forall>p\<in>#mset xs. prime p" by auto
63534
523b488b15c9 Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents: 62429
diff changeset
  1153
  ultimately have "primefact xs n" by (auto simp: primefact_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1154
  then show ?thesis ..
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
lemma primefact_contains:
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1158
  fixes p :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1159
  assumes pf: "primefact ps n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1160
    and p: "prime p"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1161
    and pn: "p dvd n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1162
  shows "p \<in> set ps"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
  using pf p pn
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1164
proof (induct ps arbitrary: p n)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1165
  case Nil
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1166
  then show ?case by (auto simp: primefact_def)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
next
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1168
  case (Cons q qs)
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
  from Cons.prems[unfolded primefact_def]
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1170
  have q: "prime q" "q * foldr (*) qs 1 = n" "\<forall>p \<in>set qs. prime p"
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1171
    and p: "prime p" "p dvd q * foldr (*) qs 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1172
    by simp_all
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1173
  consider "p dvd q" | "p dvd foldr (*) qs 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1174
    by (metis p prime_dvd_mult_eq_nat)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1175
  then show ?case
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1176
  proof cases
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1177
    case 1
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1178
    with p(1) q(1) have "p = q"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1179
      unfolding prime_nat_iff by auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1180
    then show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1181
  next
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1182
    case prem: 2
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1183
    from q(3) have pqs: "primefact qs (foldr (*) qs 1)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
      by (simp add: primefact_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1185
    from Cons.hyps[OF pqs p(1) prem] show ?thesis by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1186
  qed
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1189
lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr (*) ps 1 = n \<and> list_all prime ps"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
  by (auto simp add: primefact_def list_all_iff)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1192
text \<open>Variant of Lucas theorem.\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
lemma lucas_primefact:
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1195
    and psn: "foldr (*) ps 1 = n - 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1196
    and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
  shows "prime n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1198
proof -
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1199
  {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1200
    fix p
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1201
    assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
    from psn psp have psn1: "primefact ps (n - 1)"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
      by (auto simp add: list_all_iff primefact_variant)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
    from p(3) primefact_contains[OF psn1 p(1,2)] psp
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1205
    have False by (induct ps) auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1206
  }
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
  with lucas[OF n an] show ?thesis by blast
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1210
text \<open>Variant of Pocklington theorem.\<close>
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
lemma pocklington_primefact:
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
  assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q\<^sup>2"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68707
diff changeset
  1213
    and arnb: "(a^r) mod n = b" and psq: "foldr (*) ps 1 = q"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1214
    and bqn: "(b^q) mod n = 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1215
    and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
  shows "prime n"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1217
proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
  from bqn psp qrn
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
  have bqn: "a ^ (n - 1) mod n = 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1220
    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1221
    unfolding arnb[symmetric] power_mod
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
    by (simp_all add: power_mult[symmetric] algebra_simps)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1223
  from n have n0: "n > 0" by arith
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 63905
diff changeset
  1224
  from div_mult_mod_eq[of "a^(n - 1)" n]
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
    mod_less_divisor[OF n0, of "a^(n - 1)"]
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
  have an1: "[a ^ (n - 1) = 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1227
    by (metis bqn cong_def mod_mod_trivial)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1228
  have "coprime (a ^ ((n - 1) div p) - 1) n" if p: "prime p" "p dvd q" for p
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1229
  proof -
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
    from psp psq have pfpsq: "primefact ps q"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
      by (auto simp add: primefact_variant list_all_iff)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
    from psp primefact_contains[OF pfpsq p]
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
    have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
      by (simp add: list_all_iff)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1235
    from p prime_nat_iff have p01: "p \<noteq> 0" "p \<noteq> 1" "p = Suc (p - 1)"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
      by auto
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
    from div_mult1_eq[of r q p] p(2)
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
    have eq1: "r* (q div p) = (n - 1) div p"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55370
diff changeset
  1239
      unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1240
    have ath: "a \<le> b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> 1 \<le> a \<and> 1 \<le> b" for a b :: nat
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1241
      by arith
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1242
    {
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1243
      assume "a ^ ((n - 1) div p) mod n = 0"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1244
      then obtain s where s: "a ^ ((n - 1) div p) = n * s"
68157
057d5b4ce47e removed some non-essential rules
haftmann
parents: 67399
diff changeset
  1245
        by blast
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1246
      then have eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1247
      from qrn[symmetric] have qn1: "q dvd n - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1248
        by (auto simp: dvd_def)
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1249
      from dvd_trans[OF p(2) qn1] have npp: "(n - 1) div p * p = n - 1"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1250
        by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1251
      with eq0 have "a ^ (n - 1) = (n * s) ^ p"
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
        by (simp add: power_mult[symmetric])
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1253
      with bqn p01 have "1 = (n * s)^(Suc (p - 1)) mod n"
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1254
        by simp
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55370
diff changeset
  1255
      also have "\<dots> = 0" by (simp add: mult.assoc)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1256
      finally have False by simp
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1257
    }
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1258
    then have *: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1259
    have "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1260
      by (simp add: cong_def)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1261
    with ath[OF mod_less_eq_dividend *]
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1262
    have "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66305
diff changeset
  1263
      by (simp add: cong_diff_nat)
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1264
    then show ?thesis
68707
5b269897df9d de-applying and removal of obsolete aliases
paulson <lp15@cam.ac.uk>
parents: 68157
diff changeset
  1265
      by (metis cong_imp_coprime eq1 p')
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1266
  qed
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1267
  with pocklington[OF n qrn[symmetric] nq2 an1] show ?thesis
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
  1268
    by blast
55321
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
qed
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
eadea363deb6 Restoration of Pocklington.thy. Tidying.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
end