src/HOL/Number_Theory/Totient.thy
author nipkow
Mon, 24 Sep 2018 14:30:09 +0200
changeset 69064 5840724b1d71
parent 67399 eab6ce8368fa
child 69785 9e326f6f8a24
permissions -rw-r--r--
Prefix form of infix with * on either side no longer needs special treatment because (* and *) are no longer comment brackets in terms.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Number_Theory/Totient.thy
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     2
    Author:     Jeremy Avigad
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
     4
    Author:     Manuel Eberl
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     5
*)
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     6
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     7
section \<open>Fundamental facts about Euler's totient function\<close>
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     8
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
     9
theory Totient
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    10
imports
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    11
  Complex_Main
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66305
diff changeset
    12
  "HOL-Computational_Algebra.Primes"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    13
  "~~/src/HOL/Number_Theory/Cong"
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    14
begin
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    15
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    16
definition totatives :: "nat \<Rightarrow> nat set" where
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    17
  "totatives n = {k \<in> {0<..n}. coprime k n}"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    18
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    19
lemma in_totatives_iff: "k \<in> totatives n \<longleftrightarrow> k > 0 \<and> k \<le> n \<and> coprime k n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    20
  by (simp add: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    21
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    22
lemma totatives_code [code]: "totatives n = Set.filter (\<lambda>k. coprime k n) {0<..n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    23
  by (simp add: totatives_def Set.filter_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    24
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    25
lemma finite_totatives [simp]: "finite (totatives n)"
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    26
  by (simp add: totatives_def)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    27
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    28
lemma totatives_subset: "totatives n \<subseteq> {0<..n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    29
  by (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    30
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    31
lemma zero_not_in_totatives [simp]: "0 \<notin> totatives n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    32
  by (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    33
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    34
lemma totatives_le: "x \<in> totatives n \<Longrightarrow> x \<le> n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    35
  by (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    36
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    37
lemma totatives_less: 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    38
  assumes "x \<in> totatives n" "n > 1"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    39
  shows   "x < n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    40
proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    41
  from assms have "x \<noteq> n" by (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    42
  with totatives_le[OF assms(1)] show ?thesis by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    43
qed
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    44
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    45
lemma totatives_0 [simp]: "totatives 0 = {}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    46
  by (auto simp: totatives_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    47
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    48
lemma totatives_1 [simp]: "totatives 1 = {Suc 0}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    49
  by (auto simp: totatives_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    50
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    51
lemma totatives_Suc_0 [simp]: "totatives (Suc 0) = {Suc 0}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    52
  by (auto simp: totatives_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    53
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    54
lemma one_in_totatives [simp]: "n > 0 \<Longrightarrow> Suc 0 \<in> totatives n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    55
  by (auto simp: totatives_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    56
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    57
lemma totatives_eq_empty_iff [simp]: "totatives n = {} \<longleftrightarrow> n = 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    58
  using one_in_totatives[of n] by (auto simp del: one_in_totatives)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    59
    
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    60
lemma minus_one_in_totatives:
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    61
  assumes "n \<ge> 2"
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    62
  shows "n - 1 \<in> totatives n"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    63
  using assms coprime_diff_one_left_nat [of n] by (simp add: in_totatives_iff)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    64
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    65
lemma totatives_prime_power_Suc:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    66
  assumes "prime p"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    67
  shows   "totatives (p ^ Suc n) = {0<..p^Suc n} - (\<lambda>m. p * m) ` {0<..p^n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    68
proof safe
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    69
  fix m assume m: "p * m \<in> totatives (p ^ Suc n)" and m: "m \<in> {0<..p^n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    70
  thus False using assms by (auto simp: totatives_def gcd_mult_left)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    71
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    72
  fix k assume k: "k \<in> {0<..p^Suc n}" "k \<notin> (\<lambda>m. p * m) ` {0<..p^n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    73
  from k have "\<not>(p dvd k)" by (auto elim!: dvdE)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    74
  hence "coprime k (p ^ Suc n)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    75
    using prime_imp_coprime [OF assms, of k]
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
    76
    by (cases "n > 0") (auto simp add: ac_simps)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    77
  with k show "k \<in> totatives (p ^ Suc n)" by (simp add: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    78
qed (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    79
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    80
lemma totatives_prime: "prime p \<Longrightarrow> totatives p = {0<..<p}"
67118
ccab07d1196c more simplification rules
haftmann
parents: 67051
diff changeset
    81
  using totatives_prime_power_Suc [of p 0] by auto
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    82
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    83
lemma bij_betw_totatives:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    84
  assumes "m1 > 1" "m2 > 1" "coprime m1 m2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    85
  shows   "bij_betw (\<lambda>x. (x mod m1, x mod m2)) (totatives (m1 * m2)) 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    86
             (totatives m1 \<times> totatives m2)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    87
  unfolding bij_betw_def
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
    88
proof
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    89
  show "inj_on (\<lambda>x. (x mod m1, x mod m2)) (totatives (m1 * m2))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    90
  proof (intro inj_onI, clarify)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    91
    fix x y assume xy: "x \<in> totatives (m1 * m2)" "y \<in> totatives (m1 * m2)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    92
                       "x mod m1 = y mod m1" "x mod m2 = y mod m2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    93
    have ex: "\<exists>!z. z < m1 * m2 \<and> [z = x] (mod m1) \<and> [z = x] (mod m2)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    94
      by (rule binary_chinese_remainder_unique_nat) (insert assms, simp_all)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    95
    have "x < m1 * m2 \<and> [x = x] (mod m1) \<and> [x = x] (mod m2)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    96
         "y < m1 * m2 \<and> [y = x] (mod m1) \<and> [y = x] (mod m2)"
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66803
diff changeset
    97
      using xy assms by (simp_all add: totatives_less one_less_mult cong_def)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    98
    from this[THEN the1_equality[OF ex]] show "x = y" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
    99
  qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   100
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   101
  show "(\<lambda>x. (x mod m1, x mod m2)) ` totatives (m1 * m2) = totatives m1 \<times> totatives m2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   102
  proof safe
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   103
    fix x assume "x \<in> totatives (m1 * m2)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   104
    with assms show "x mod m1 \<in> totatives m1" "x mod m2 \<in> totatives m2"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   105
      using coprime_common_divisor [of x m1 m1] coprime_common_divisor [of x m2 m2]
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   106
      by (auto simp add: in_totatives_iff mod_greater_zero_iff_not_dvd)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   107
  next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   108
    fix a b assume ab: "a \<in> totatives m1" "b \<in> totatives m2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   109
    with assms have ab': "a < m1" "b < m2" by (auto simp: totatives_less)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   110
    with binary_chinese_remainder_unique_nat[OF assms(3), of a b] obtain x
66888
930abfdf8727 algebraic foundation for congruences
haftmann
parents: 66803
diff changeset
   111
      where x: "x < m1 * m2" "x mod m1 = a" "x mod m2 = b" by (auto simp: cong_def)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   112
    from x ab assms(3) have "x \<in> totatives (m1 * m2)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   113
      by (auto intro: ccontr simp add: in_totatives_iff)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   114
    with x show "(a, b) \<in> (\<lambda>x. (x mod m1, x mod m2)) ` totatives (m1*m2)" by blast
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   115
  qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   116
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   117
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   118
lemma bij_betw_totatives_gcd_eq:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   119
  fixes n d :: nat
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   120
  assumes "d dvd n" "n > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   121
  shows   "bij_betw (\<lambda>k. k * d) (totatives (n div d)) {k\<in>{0<..n}. gcd k n = d}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   122
  unfolding bij_betw_def
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   123
proof
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   124
  show "inj_on (\<lambda>k. k * d) (totatives (n div d))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   125
    by (auto simp: inj_on_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   126
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   127
  show "(\<lambda>k. k * d) ` totatives (n div d) = {k\<in>{0<..n}. gcd k n = d}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   128
  proof (intro equalityI subsetI, goal_cases)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   129
    case (1 k)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   130
    then show ?case using assms
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   131
      by (auto elim: dvdE simp add: in_totatives_iff ac_simps gcd_mult_right)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   132
  next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   133
    case (2 k)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   134
    hence "d dvd k" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   135
    then obtain l where k: "k = l * d" by (elim dvdE) auto
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   136
    from 2 assms show ?case
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   137
      using gcd_mult_right [of _ d l]
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   138
      by (auto intro: gcd_eq_1_imp_coprime elim!: dvdE simp add: k image_iff in_totatives_iff ac_simps)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   139
  qed
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   140
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   141
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   142
definition totient :: "nat \<Rightarrow> nat" where
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   143
  "totient n = card (totatives n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   144
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   145
primrec totient_naive :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   146
  "totient_naive 0 acc n = acc"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   147
| "totient_naive (Suc k) acc n =
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   148
     (if coprime (Suc k) n then totient_naive k (acc + 1) n else totient_naive k acc n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   149
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   150
lemma totient_naive:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   151
  "totient_naive k acc n = card {x \<in> {0<..k}. coprime x n} + acc"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   152
proof (induction k arbitrary: acc)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   153
  case (Suc k acc)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   154
  have "totient_naive (Suc k) acc n = 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   155
          (if coprime (Suc k) n then 1 else 0) + card {x \<in> {0<..k}. coprime x n} + acc"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   156
    using Suc by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   157
  also have "(if coprime (Suc k) n then 1 else 0) = 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   158
               card (if coprime (Suc k) n then {Suc k} else {})" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   159
  also have "\<dots> + card {x \<in> {0<..k}. coprime x n} =
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   160
               card ((if coprime (Suc k) n then {Suc k} else {}) \<union> {x \<in> {0<..k}. coprime x n})"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   161
    by (intro card_Un_disjoint [symmetric]) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   162
  also have "((if coprime (Suc k) n then {Suc k} else {}) \<union> {x \<in> {0<..k}. coprime x n}) =
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   163
               {x \<in> {0<..Suc k}. coprime x n}" by (auto elim: le_SucE)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   164
  finally show ?case .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   165
qed simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   166
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   167
lemma totient_code_naive [code]: "totient n = totient_naive n 0 n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   168
  by (subst totient_naive) (simp add: totient_def totatives_def)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   169
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   170
lemma totient_le: "totient n \<le> n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   171
proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   172
  have "card (totatives n) \<le> card {0<..n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   173
    by (intro card_mono) (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   174
  thus ?thesis by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   175
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   176
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   177
lemma totient_less: 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   178
  assumes "n > 1"
66305
7454317f883c misc tuning and modernization;
wenzelm
parents: 65726
diff changeset
   179
  shows "totient n < n"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   180
proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   181
  from assms have "card (totatives n) \<le> card {0<..<n}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   182
    using totatives_less[of _ n] totatives_subset[of n] by (intro card_mono) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   183
  with assms show ?thesis by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   184
qed    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   185
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   186
lemma totient_0 [simp]: "totient 0 = 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   187
  by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   188
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   189
lemma totient_Suc_0 [simp]: "totient (Suc 0) = Suc 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   190
  by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   191
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   192
lemma totient_1 [simp]: "totient 1 = Suc 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   193
  by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   194
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   195
lemma totient_0_iff [simp]: "totient n = 0 \<longleftrightarrow> n = 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   196
  by (auto simp: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   197
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   198
lemma totient_gt_0_iff [simp]: "totient n > 0 \<longleftrightarrow> n > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   199
  by (auto intro: Nat.gr0I)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   200
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   201
lemma card_gcd_eq_totient:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   202
  "n > 0 \<Longrightarrow> d dvd n \<Longrightarrow> card {k\<in>{0<..n}. gcd k n = d} = totient (n div d)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   203
  unfolding totient_def by (rule sym, rule bij_betw_same_card[OF bij_betw_totatives_gcd_eq])
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   204
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   205
lemma totient_divisor_sum: "(\<Sum>d | d dvd n. totient d) = n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   206
proof (cases "n = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   207
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   208
  hence "n > 0" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   209
  define A where "A = (\<lambda>d. {k\<in>{0<..n}. gcd k n = d})"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   210
  have *: "card (A d) = totient (n div d)" if d: "d dvd n" for d
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   211
    using \<open>n > 0\<close> and d unfolding A_def by (rule card_gcd_eq_totient)  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   212
  have "n = card {1..n}" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   213
  also have "{1..n} = (\<Union>d\<in>{d. d dvd n}. A d)" by safe (auto simp: A_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   214
  also have "card \<dots> = (\<Sum>d | d dvd n. card (A d))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   215
    using \<open>n > 0\<close> by (intro card_UN_disjoint) (auto simp: A_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   216
  also have "\<dots> = (\<Sum>d | d dvd n. totient (n div d))" by (intro sum.cong refl *) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   217
  also have "\<dots> = (\<Sum>d | d dvd n. totient d)" using \<open>n > 0\<close>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   218
    by (intro sum.reindex_bij_witness[of _ "(div) n" "(div) n"]) (auto elim: dvdE)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   219
  finally show ?thesis ..
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   220
qed auto
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   221
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   222
lemma totient_mult_coprime:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   223
  assumes "coprime m n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   224
  shows   "totient (m * n) = totient m * totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   225
proof (cases "m > 1 \<and> n > 1")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   226
  case True
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   227
  hence mn: "m > 1" "n > 1" by simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   228
  have "totient (m * n) = card (totatives (m * n))" by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   229
  also have "\<dots> = card (totatives m \<times> totatives n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   230
    using bij_betw_totatives [OF mn \<open>coprime m n\<close>] by (rule bij_betw_same_card)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   231
  also have "\<dots> = totient m * totient n" by (simp add: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   232
  finally show ?thesis .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   233
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   234
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   235
  with assms show ?thesis by (cases m; cases n) auto
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   236
qed
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   237
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   238
lemma totient_prime_power_Suc:
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   239
  assumes "prime p"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   240
  shows   "totient (p ^ Suc n) = p ^ n * (p - 1)"
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   241
proof -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67399
diff changeset
   242
  from assms have "totient (p ^ Suc n) = card ({0<..p ^ Suc n} - (*) p ` {0<..p ^ n})"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   243
    unfolding totient_def by (subst totatives_prime_power_Suc) simp_all
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67399
diff changeset
   244
  also from assms have "\<dots> = p ^ Suc n - card ((*) p ` {0<..p^n})"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   245
    by (subst card_Diff_subset) (auto intro: prime_gt_0_nat)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67399
diff changeset
   246
  also from assms have "card ((*) p ` {0<..p^n}) = p ^ n"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   247
    by (subst card_image) (auto simp: inj_on_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   248
  also have "p ^ Suc n - p ^ n = p ^ n * (p - 1)" by (simp add: algebra_simps)
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   249
  finally show ?thesis .
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   250
qed
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   251
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   252
lemma totient_prime_power:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   253
  assumes "prime p" "n > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   254
  shows   "totient (p ^ n) = p ^ (n - 1) * (p - 1)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   255
  using totient_prime_power_Suc[of p "n - 1"] assms by simp
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   256
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   257
lemma totient_imp_prime:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   258
  assumes "totient p = p - 1" "p > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   259
  shows   "prime p"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   260
proof (cases "p = 1")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   261
  case True
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   262
  with assms show ?thesis by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   263
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   264
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   265
  with assms have p: "p > 1" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   266
  have "x \<in> {0<..<p}" if "x \<in> totatives p" for x
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   267
    using that and p by (cases "x = p") (auto simp: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   268
  with assms have *: "totatives p = {0<..<p}"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   269
    by (intro card_subset_eq) (auto simp: totient_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   270
  have **: False if "x \<noteq> 1" "x \<noteq> p" "x dvd p" for x
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   271
  proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   272
    from that have nz: "x \<noteq> 0" by (auto intro!: Nat.gr0I)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   273
    from that and p have le: "x \<le> p" by (intro dvd_imp_le) auto
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   274
    from that and nz have "\<not>coprime x p"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   275
      by (auto elim: dvdE)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   276
    hence "x \<notin> totatives p" by (simp add: totatives_def)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   277
    also note *
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   278
    finally show False using that and le by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   279
  qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   280
  hence "(\<forall>m. m dvd p \<longrightarrow> m = 1 \<or> m = p)" by blast
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   281
  with p show ?thesis by (subst prime_nat_iff) (auto dest: **)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   282
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   283
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   284
lemma totient_prime:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   285
  assumes "prime p"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   286
  shows   "totient p = p - 1"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   287
  using totient_prime_power_Suc[of p 0] assms by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   288
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   289
lemma totient_2 [simp]: "totient 2 = 1"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   290
  and totient_3 [simp]: "totient 3 = 2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   291
  and totient_5 [simp]: "totient 5 = 4"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   292
  and totient_7 [simp]: "totient 7 = 6"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   293
  by (subst totient_prime; simp)+
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   294
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   295
lemma totient_4 [simp]: "totient 4 = 2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   296
  and totient_8 [simp]: "totient 8 = 4"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   297
  and totient_9 [simp]: "totient 9 = 6"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   298
  using totient_prime_power[of 2 2] totient_prime_power[of 2 3] totient_prime_power[of 3 2] 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   299
  by simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   300
    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   301
lemma totient_6 [simp]: "totient 6 = 2"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   302
  using totient_mult_coprime [of 2 3] coprime_add_one_right [of 2]
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   303
  by simp
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   304
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   305
lemma totient_even:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   306
  assumes "n > 2"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   307
  shows   "even (totient n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   308
proof (cases "\<exists>p. prime p \<and> p \<noteq> 2 \<and> p dvd n")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   309
  case True
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   310
  then obtain p where p: "prime p" "p \<noteq> 2" "p dvd n" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   311
  from \<open>p \<noteq> 2\<close> have "p = 0 \<or> p = 1 \<or> p > 2" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   312
  with p(1) have "odd p" using prime_odd_nat[of p] by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   313
  define k where "k = multiplicity p n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   314
  from p assms have k_pos: "k > 0" unfolding k_def by (subst multiplicity_gt_zero_iff) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   315
  have "p ^ k dvd n" unfolding k_def by (simp add: multiplicity_dvd)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   316
  then obtain m where m: "n = p ^ k * m" by (elim dvdE)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   317
  with assms have m_pos: "m > 0" by (auto intro!: Nat.gr0I)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   318
  from k_def m_pos p have "\<not> p dvd m"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   319
    by (subst (asm) m) (auto intro!: Nat.gr0I simp: prime_elem_multiplicity_mult_distrib 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   320
                          prime_elem_multiplicity_eq_zero_iff)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   321
  with \<open>prime p\<close> have "coprime p m"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   322
    by (rule prime_imp_coprime)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   323
  with \<open>k > 0\<close> have "coprime (p ^ k) m"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   324
    by simp
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   325
  then show ?thesis using p k_pos \<open>odd p\<close> 
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   326
    by (auto simp add: m totient_mult_coprime totient_prime_power)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   327
next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   328
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   329
  from assms have "n = (\<Prod>p\<in>prime_factors n. p ^ multiplicity p n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   330
    by (intro Primes.prime_factorization_nat) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   331
  also from False have "\<dots> = (\<Prod>p\<in>prime_factors n. if p = 2 then 2 ^ multiplicity 2 n else 1)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   332
    by (intro prod.cong refl) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   333
  also have "\<dots> = 2 ^ multiplicity 2 n" 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   334
    by (subst prod.delta[OF finite_set_mset]) (auto simp: prime_factors_multiplicity)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   335
  finally have n: "n = 2 ^ multiplicity 2 n" .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   336
  have "multiplicity 2 n = 0 \<or> multiplicity 2 n = 1 \<or> multiplicity 2 n > 1" by force
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   337
  with n assms have "multiplicity 2 n > 1" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   338
  thus ?thesis by (subst n) (simp add: totient_prime_power)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   339
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   340
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   341
lemma totient_prod_coprime:
66803
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   342
  assumes "pairwise coprime (f ` A)" "inj_on f A"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   343
  shows   "totient (prod f A) = (\<Prod>a\<in>A. totient (f a))"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   344
  using assms
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   345
proof (induction A rule: infinite_finite_induct)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   346
  case (insert x A)
66803
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   347
  have *: "coprime (prod f A) (f x)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   348
  proof (rule prod_coprime_left)
66803
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   349
    fix y
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   350
    assume "y \<in> A"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   351
    with \<open>x \<notin> A\<close> have "y \<noteq> x"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   352
      by auto
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   353
    with \<open>x \<notin> A\<close> \<open>y \<in> A\<close> \<open>inj_on f (insert x A)\<close> have "f y \<noteq> f x"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   354
      using inj_onD [of f "insert x A" y x]
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   355
      by auto
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   356
    with \<open>y \<in> A\<close> show "coprime (f y) (f x)"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   357
      using pairwiseD [OF \<open>pairwise coprime (f ` insert x A)\<close>]
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   358
      by auto
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   359
  qed
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   360
  from insert.hyps have "prod f (insert x A) = prod f A * f x" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   361
  also have "totient \<dots> = totient (prod f A) * totient (f x)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   362
    using insert.hyps insert.prems by (intro totient_mult_coprime *)
66803
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   363
  also have "totient (prod f A) = (\<Prod>a\<in>A. totient (f a))" 
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   364
    using insert.prems by (intro insert.IH) (auto dest: pairwise_subset)
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   365
  also from insert.hyps have "\<dots> * totient (f x) = (\<Prod>a\<in>insert x A. totient (f a))" by simp
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   366
  finally show ?case .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   367
qed simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   368
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   369
(* TODO Move *)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   370
lemma prime_power_eq_imp_eq:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   371
  fixes p q :: "'a :: factorial_semiring"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   372
  assumes "prime p" "prime q" "m > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   373
  assumes "p ^ m = q ^ n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   374
  shows   "p = q"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   375
proof (rule ccontr)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   376
  assume pq: "p \<noteq> q"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   377
  from assms have "m = multiplicity p (p ^ m)" 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   378
    by (subst multiplicity_prime_power) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   379
  also note \<open>p ^ m = q ^ n\<close>
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   380
  also from assms pq have "multiplicity p (q ^ n) = 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   381
    by (subst multiplicity_distinct_prime_power) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   382
  finally show False using \<open>m > 0\<close> by simp
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   383
qed
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   384
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   385
lemma totient_formula1:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   386
  assumes "n > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   387
  shows   "totient n = (\<Prod>p\<in>prime_factors n. p ^ (multiplicity p n - 1) * (p - 1))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   388
proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   389
  from assms have "n = (\<Prod>p\<in>prime_factors n. p ^ multiplicity p n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   390
    by (rule prime_factorization_nat)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   391
  also have "totient \<dots> = (\<Prod>x\<in>prime_factors n. totient (x ^ multiplicity x n))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   392
  proof (rule totient_prod_coprime)
66803
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   393
    show "pairwise coprime ((\<lambda>p. p ^ multiplicity p n) ` prime_factors n)"
dd8922885a68 avoid trivial definition
haftmann
parents: 66453
diff changeset
   394
    proof (rule pairwiseI, clarify)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   395
      fix p q assume *: "p \<in># prime_factorization n" "q \<in># prime_factorization n" 
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   396
                     "p ^ multiplicity p n \<noteq> q ^ multiplicity q n"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   397
      then have "multiplicity p n > 0" "multiplicity q n > 0"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   398
        by (simp_all add: prime_factors_multiplicity)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   399
      with * primes_coprime [of p q] show "coprime (p ^ multiplicity p n) (q ^ multiplicity q n)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   400
        by auto
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   401
    qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   402
  next
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   403
    show "inj_on (\<lambda>p. p ^ multiplicity p n) (prime_factors n)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   404
    proof
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   405
      fix p q assume pq: "p \<in># prime_factorization n" "q \<in># prime_factorization n" 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   406
                         "p ^ multiplicity p n = q ^ multiplicity q n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   407
      from assms and pq have "prime p" "prime q" "multiplicity p n > 0" 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   408
        by (simp_all add: prime_factors_multiplicity)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   409
      from prime_power_eq_imp_eq[OF this pq(3)] show "p = q" .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   410
    qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   411
  qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   412
  also have "\<dots> = (\<Prod>p\<in>prime_factors n. p ^ (multiplicity p n - 1) * (p - 1))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   413
    by (intro prod.cong refl totient_prime_power) (auto simp: prime_factors_multiplicity)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   414
  finally show ?thesis .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   415
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   416
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   417
lemma totient_dvd:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   418
  assumes "m dvd n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   419
  shows   "totient m dvd totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   420
proof (cases "m = 0 \<or> n = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   421
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   422
  let ?M = "\<lambda>p m :: nat. multiplicity p m - 1"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   423
  have "(\<Prod>p\<in>prime_factors m. p ^ ?M p m * (p - 1)) dvd
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   424
          (\<Prod>p\<in>prime_factors n. p ^ ?M p n * (p - 1))" using assms False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   425
    by (intro prod_dvd_prod_subset2 mult_dvd_mono dvd_refl le_imp_power_dvd diff_le_mono
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   426
              dvd_prime_factors dvd_imp_multiplicity_le) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   427
  with False show ?thesis by (simp add: totient_formula1)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   428
qed (insert assms, auto)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   429
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   430
lemma totient_dvd_mono:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   431
  assumes "m dvd n" "n > 0"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   432
  shows   "totient m \<le> totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   433
  by (cases "m = 0") (insert assms, auto intro: dvd_imp_le totient_dvd)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   434
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   435
(* TODO Move *)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   436
lemma prime_factors_power: "n > 0 \<Longrightarrow> prime_factors (x ^ n) = prime_factors x"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   437
  by (cases "x = 0"; cases "n = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   438
     (auto simp: prime_factors_multiplicity prime_elem_multiplicity_power_distrib zero_power)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   439
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   440
lemma totient_formula2:
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   441
  "real (totient n) = real n * (\<Prod>p\<in>prime_factors n. 1 - 1 / real p)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   442
proof (cases "n = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   443
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   444
  have "real (totient n) = (\<Prod>p\<in>prime_factors n. real
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   445
          (p ^ (multiplicity p n - 1) * (p - 1)))" 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   446
    using False by (subst totient_formula1) simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   447
  also have "\<dots> = (\<Prod>p\<in>prime_factors n. real (p ^ multiplicity p n) * (1 - 1 / real p))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   448
    by (intro prod.cong refl) (auto simp add: field_simps prime_factors_multiplicity 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   449
          prime_ge_Suc_0_nat of_nat_diff power_Suc [symmetric] simp del: power_Suc)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   450
  also have "\<dots> = real (\<Prod>p\<in>prime_factors n. p ^ multiplicity p n) * 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   451
                    (\<Prod>p\<in>prime_factors n. 1 - 1 / real p)" by (subst prod.distrib) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   452
  also have "(\<Prod>p\<in>prime_factors n. p ^ multiplicity p n) = n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   453
    using False by (intro Primes.prime_factorization_nat [symmetric]) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   454
  finally show ?thesis .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   455
qed auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   456
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   457
lemma totient_gcd: "totient (a * b) * totient (gcd a b) = totient a * totient b * gcd a b"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   458
proof (cases "a = 0 \<or> b = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   459
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   460
  let ?P = "prime_factors :: nat \<Rightarrow> nat set"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   461
  have "real (totient a * totient b * gcd a b) = real (a * b * gcd a b) *
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   462
          ((\<Prod>p\<in>?P a. 1 - 1 / real p) * (\<Prod>p\<in>?P b. 1 - 1 / real p))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   463
    by (simp add: totient_formula2)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   464
  also have "?P a = (?P a - ?P b) \<union> (?P a \<inter> ?P b)" by auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   465
  also have "(\<Prod>p\<in>\<dots>. 1 - 1 / real p) = 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   466
                 (\<Prod>p\<in>?P a - ?P b. 1 - 1 / real p) * (\<Prod>p\<in>?P a \<inter> ?P b. 1 - 1 / real p)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   467
    by (rule prod.union_disjoint) blast+
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   468
  also have "\<dots> * (\<Prod>p\<in>?P b. 1 - 1 / real p) = (\<Prod>p\<in>?P a - ?P b. 1 - 1 / real p) * 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   469
               (\<Prod>p\<in>?P b. 1 - 1 / real p) * (\<Prod>p\<in>?P a \<inter> ?P b. 1 - 1 / real p)" (is "_ = ?A * _")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   470
    by (simp only: mult_ac)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   471
  also have "?A = (\<Prod>p\<in>?P a - ?P b \<union> ?P b. 1 - 1 / real p)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   472
    by (rule prod.union_disjoint [symmetric]) blast+
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   473
  also have "?P a - ?P b \<union> ?P b = ?P a \<union> ?P b" by blast
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   474
  also have "real (a * b * gcd a b) * ((\<Prod>p\<in>\<dots>. 1 - 1 / real p) * 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   475
                 (\<Prod>p\<in>?P a \<inter> ?P b. 1 - 1 / real p)) = real (totient (a * b) * totient (gcd a b))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   476
    using False by (simp add: totient_formula2 prime_factors_product prime_factorization_gcd)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   477
  finally show ?thesis by (simp only: of_nat_eq_iff)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   478
qed auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   479
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   480
lemma totient_mult: "totient (a * b) = totient a * totient b * gcd a b div totient (gcd a b)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   481
  by (subst totient_gcd [symmetric]) simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   482
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   483
lemma of_nat_eq_1_iff: "of_nat x = (1 :: 'a :: {semiring_1, semiring_char_0}) \<longleftrightarrow> x = 1"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   484
  by (fact of_nat_eq_1_iff)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   485
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   486
(* TODO Move *)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   487
lemma odd_imp_coprime_nat:
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   488
  assumes "odd (n::nat)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   489
  shows   "coprime n 2"
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   490
proof -
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   491
  from assms obtain k where n: "n = Suc (2 * k)" by (auto elim!: oddE)
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   492
  have "coprime (Suc (2 * k)) (2 * k)"
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   493
    by (fact coprime_Suc_left_nat)
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   494
  then show ?thesis using n
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   495
    by simp
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   496
qed
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   497
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   498
lemma totient_double: "totient (2 * n) = (if even n then 2 * totient n else totient n)"
67051
e7e54a0b9197 dedicated definition for coprimality
haftmann
parents: 66888
diff changeset
   499
  by (simp add: totient_mult ac_simps odd_imp_coprime_nat)
65726
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   500
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   501
lemma totient_power_Suc: "totient (n ^ Suc m) = n ^ m * totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   502
proof (induction m arbitrary: n)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   503
  case (Suc m n)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   504
  have "totient (n ^ Suc (Suc m)) = totient (n * n ^ Suc m)" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   505
  also have "\<dots> = n ^ Suc m * totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   506
    using Suc.IH by (subst totient_mult) simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   507
  finally show ?case .
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   508
qed simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   509
  
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   510
lemma totient_power: "m > 0 \<Longrightarrow> totient (n ^ m) = n ^ (m - 1) * totient n"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   511
  using totient_power_Suc[of n "m - 1"] by (cases m) simp_all
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   512
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   513
lemma totient_gcd_lcm: "totient (gcd a b) * totient (lcm a b) = totient a * totient b"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   514
proof (cases "a = 0 \<or> b = 0")
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   515
  case False
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   516
  let ?P = "prime_factors :: nat \<Rightarrow> nat set" and ?f = "\<lambda>p::nat. 1 - 1 / real p"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   517
  have "real (totient (gcd a b) * totient (lcm a b)) = real (gcd a b * lcm a b) * 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   518
          (prod ?f (?P a \<inter> ?P b) * prod ?f (?P a \<union> ?P b))"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   519
    using False unfolding of_nat_mult 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   520
    by (simp add: totient_formula2 prime_factorization_gcd prime_factorization_lcm)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   521
  also have "gcd a b * lcm a b = a * b" by simp
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   522
  also have "?P a \<union> ?P b = (?P a - ?P a \<inter> ?P b) \<union> ?P b" by blast
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   523
  also have "prod ?f \<dots> = prod ?f (?P a - ?P a \<inter> ?P b) * prod ?f (?P b)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   524
    by (rule prod.union_disjoint) blast+
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   525
  also have "prod ?f (?P a \<inter> ?P b) * \<dots> = 
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   526
               prod ?f (?P a \<inter> ?P b \<union> (?P a - ?P a \<inter> ?P b)) * prod ?f (?P b)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   527
    by (subst prod.union_disjoint) auto
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   528
  also have "?P a \<inter> ?P b \<union> (?P a - ?P a \<inter> ?P b) = ?P a" by blast
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   529
  also have "real (a * b) * (prod ?f (?P a) * prod ?f (?P b)) = real (totient a * totient b)"
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   530
    using False by (simp add: totient_formula2)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   531
  finally show ?thesis by (simp only: of_nat_eq_iff)
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   532
qed auto    
f5d64d094efe More material on totient function
eberlm <eberlm@in.tum.de>
parents: 65465
diff changeset
   533
65465
067210a08a22 more fundamental euler's totient function on nat rather than int;
haftmann
parents:
diff changeset
   534
end