src/HOL/Number_Theory/Modular_Inverse.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82518 da14e77a48b2
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     1
section \<open>Modular Inverses\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     2
theory Modular_Inverse
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     3
  imports Cong "HOL-Library.FuncSet"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     4
begin
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     5
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     6
text \<open>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     7
  The following returns the unique number $m$ such that $mn \equiv 1\pmod{p}$ if there is one,
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     8
  i.e.\ if $n$ and $p$ are coprime, and otherwise $0$ by convention.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
     9
\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    10
definition modular_inverse where
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    11
  "modular_inverse p n = (if coprime p n then fst (bezout_coefficients n p) mod p else 0)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    12
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    13
lemma cong_modular_inverse1:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    14
  assumes "coprime n p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    15
  shows   "[n * modular_inverse p n = 1] (mod p)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    16
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    17
  have "[fst (bezout_coefficients n p) * n + snd (bezout_coefficients n p) * p =
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    18
         modular_inverse p n * n + 0] (mod p)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    19
    unfolding modular_inverse_def using assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    20
    by (intro cong_add cong_mult) (auto simp: Cong.cong_def coprime_commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    21
  also have "fst (bezout_coefficients n p) * n + snd (bezout_coefficients n p) * p = gcd n p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    22
    by (simp add: bezout_coefficients_fst_snd)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    23
  also have "\<dots> = 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    24
    using assms by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    25
  finally show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    26
    by (simp add: cong_sym mult_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    27
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    28
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    29
lemma cong_modular_inverse2:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    30
  assumes "coprime n p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    31
  shows   "[modular_inverse p n * n = 1] (mod p)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    32
  using cong_modular_inverse1[OF assms] by (simp add: mult.commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    33
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    34
lemma coprime_modular_inverse [simp, intro]:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    35
  fixes n :: "'a :: {euclidean_ring_gcd,unique_euclidean_semiring}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    36
  assumes "coprime n p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    37
  shows   "coprime (modular_inverse p n) p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    38
  using cong_modular_inverse1[OF assms] assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    39
  by (meson cong_imp_coprime cong_sym coprime_1_left coprime_mult_left_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    40
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    41
lemma modular_inverse_int_nonneg: "p > 0 \<Longrightarrow> modular_inverse p (n :: int) \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    42
  by (simp add: modular_inverse_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    43
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    44
lemma modular_inverse_int_less: "p > 0 \<Longrightarrow> modular_inverse p (n :: int) < p"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    45
  by (simp add: modular_inverse_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    46
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    47
lemma modular_inverse_int_eqI:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    48
  fixes x y :: int
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    49
  assumes "y \<in> {0..<m}" "[x * y = 1] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    50
  shows   "modular_inverse m x = y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    51
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    52
  from assms have "coprime x m"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    53
    using cong_gcd_eq by force
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    54
  have "[modular_inverse m x * 1 = modular_inverse m x * (x * y)] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    55
    by (rule cong_sym, intro cong_mult assms cong_refl)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    56
  also have "modular_inverse m x * (x * y) = (modular_inverse m x * x) * y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    57
    by (simp add: mult_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    58
  also have "[\<dots> = 1 * y] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    59
    using \<open>coprime x m\<close> by (intro cong_mult cong_refl cong_modular_inverse2)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    60
  finally have "[modular_inverse m x = y] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    61
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    62
  thus "modular_inverse m x = y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    63
    using assms by (simp add: Cong.cong_def modular_inverse_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    64
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    65
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    66
lemma modular_inverse_1 [simp]:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    67
  assumes "m > (1 :: int)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    68
  shows   "modular_inverse m 1 = 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    69
  by (rule modular_inverse_int_eqI) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    70
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    71
lemma modular_inverse_int_mult:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    72
  fixes x y :: int
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    73
  assumes "coprime x m" "coprime y m" "m > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    74
  shows "modular_inverse m (x * y) = (modular_inverse m y * modular_inverse m x) mod m"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    75
proof (rule modular_inverse_int_eqI)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    76
  show "modular_inverse m y * modular_inverse m x mod m \<in> {0..<m}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    77
    using assms by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    78
next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    79
  have "[x * y * (modular_inverse m y * modular_inverse m x mod m) =
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    80
         x * y * (modular_inverse m y * modular_inverse m x)] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    81
    by (intro cong_mult cong_refl) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    82
  also have "x * y * (modular_inverse m y * modular_inverse m x) =
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    83
             (x * modular_inverse m x) * (y * modular_inverse m y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    84
    by (simp add: mult_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    85
  also have "[\<dots> = 1 * 1] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    86
    by (intro cong_mult cong_modular_inverse1 assms)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    87
  finally show "[x * y * (modular_inverse m y * modular_inverse m x mod m) = 1] (mod m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    88
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    89
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    90
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    91
lemma bij_betw_int_remainders_mult:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    92
  fixes a n :: int
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    93
  assumes a: "coprime a n"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    94
  shows   "bij_betw (\<lambda>m. a * m mod n) {1..<n} {1..<n}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    95
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    96
  define a' where "a' = modular_inverse n a"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    97
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    98
  have *: "a' * (a * m mod n) mod n = m \<and> a * m mod n \<in> {1..<n}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
    99
    if a: "[a * a' = 1] (mod n)" and m: "m \<in> {1..<n}" for m a a' :: int
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   100
  proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   101
    have "[a' * (a * m mod n) = a' * (a * m)] (mod n)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   102
      by (intro cong_mult cong_refl) (auto simp: Cong.cong_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   103
    also have "a' * (a * m) = (a * a') * m"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   104
      by (simp add: mult_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   105
    also have "[(a * a') * m = 1 * m] (mod n)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   106
      unfolding a'_def by (intro cong_mult cong_refl) (use a in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   107
    finally show "a' * (a * m mod n) mod n = m"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   108
      using m by (simp add: Cong.cong_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   109
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   110
    have "coprime a n"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   111
      using a coprime_iff_invertible_int by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   112
    hence "\<not>n dvd (a * m)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   113
      using m by (simp add: coprime_commute coprime_dvd_mult_right_iff zdvd_not_zless)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   114
    hence "a * m mod n > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   115
      using m order_le_less by fastforce
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   116
    thus "a * m mod n \<in> {1..<n}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   117
      using m by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   118
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   119
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   120
  have "[a * a' = 1] (mod n)" "[a' * a = 1] (mod n)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   121
    unfolding a'_def by (rule cong_modular_inverse1 cong_modular_inverse2; fact)+
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   122
  from this[THEN *] show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   123
    by (intro bij_betwI[of _ _ _ "\<lambda>m. a' * m mod n"]) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   124
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   125
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   126
lemma mult_modular_inverse_of_not_coprime [simp]: "\<not>coprime a m \<Longrightarrow> modular_inverse m a = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   127
  by (simp add: coprime_commute modular_inverse_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   128
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   129
lemma mult_modular_inverse_eq_0_iff:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   130
  fixes a :: "'a :: {unique_euclidean_semiring, euclidean_ring_gcd}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   131
  shows "\<not>is_unit m \<Longrightarrow> modular_inverse m a = 0 \<longleftrightarrow> \<not>coprime a m"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   132
  by (metis coprime_modular_inverse mult_modular_inverse_of_not_coprime coprime_0_left_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   133
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   134
lemma mult_modular_inverse_int_pos: "m > 1 \<Longrightarrow> coprime a m \<Longrightarrow> modular_inverse m a > (0 :: int)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   135
  by (simp add: modular_inverse_int_nonneg mult_modular_inverse_eq_0_iff order.strict_iff_order)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   136
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   137
lemma abs_mult_modular_inverse_int_less: "m \<noteq> 0 \<Longrightarrow> \<bar>modular_inverse m a :: int\<bar> < \<bar>m\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   138
  by (auto simp: modular_inverse_def intro!: abs_mod_less)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   139
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   140
lemma modular_inverse_int_less': "m \<noteq> 0 \<Longrightarrow> (modular_inverse m a :: int) < \<bar>m\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   141
  using abs_mult_modular_inverse_int_less[of m a] by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   142
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents:
diff changeset
   143
end