| author | wenzelm | 
| Thu, 16 Jan 2020 16:15:25 +0100 | |
| changeset 71388 | 57861bd0a3e1 | 
| parent 68707 | 5b269897df9d | 
| child 72302 | d7d90ed4c74e | 
| permissions | -rw-r--r-- | 
| 65413 | 1 | (* Title: HOL/Number_Theory/Quadratic_Reciprocity.thy | 
| 2 | Author: Jaime Mendizabal Roche | |
| 3 | *) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 4 | |
| 64318 
1e92b5c35615
Repaired LaTeX in HOL-Data_Structures
 eberlm <eberlm@in.tum.de> parents: 
64282diff
changeset | 5 | theory Quadratic_Reciprocity | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 6 | imports Gauss | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 7 | begin | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 8 | |
| 65413 | 9 | text \<open> | 
| 10 | The proof is based on Gauss's fifth proof, which can be found at | |
| 68224 | 11 | \<^url>\<open>https://www.lehigh.edu/~shw2/q-recip/gauss5.pdf\<close>. | 
| 65413 | 12 | \<close> | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 13 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 14 | locale QR = | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 15 | fixes p :: "nat" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 16 | fixes q :: "nat" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 17 | assumes p_prime: "prime p" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 18 | assumes p_ge_2: "2 < p" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 19 | assumes q_prime: "prime q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 20 | assumes q_ge_2: "2 < q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 21 | assumes pq_neq: "p \<noteq> q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 22 | begin | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 23 | |
| 65413 | 24 | lemma odd_p: "odd p" | 
| 25 | using p_ge_2 p_prime prime_odd_nat by blast | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 26 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 27 | lemma p_ge_0: "0 < int p" | 
| 67118 | 28 | by (simp add: p_prime prime_gt_0_nat) | 
| 29 | ||
| 65413 | 30 | lemma p_eq2: "int p = (2 * ((int p - 1) div 2)) + 1" | 
| 31 | using odd_p by simp | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 32 | |
| 65413 | 33 | lemma odd_q: "odd q" | 
| 34 | using q_ge_2 q_prime prime_odd_nat by blast | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 35 | |
| 65413 | 36 | lemma q_ge_0: "0 < int q" | 
| 67118 | 37 | by (simp add: q_prime prime_gt_0_nat) | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 38 | |
| 65413 | 39 | lemma q_eq2: "int q = (2 * ((int q - 1) div 2)) + 1" | 
| 40 | using odd_q by simp | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 41 | |
| 65413 | 42 | lemma pq_eq2: "int p * int q = (2 * ((int p * int q - 1) div 2)) + 1" | 
| 43 | using odd_p odd_q by simp | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 44 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 45 | lemma pq_coprime: "coprime p q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 46 | using pq_neq p_prime primes_coprime_nat q_prime by blast | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 47 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 48 | lemma pq_coprime_int: "coprime (int p) (int q)" | 
| 66837 | 49 | by (simp add: gcd_int_def pq_coprime) | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 50 | |
| 65413 | 51 | lemma qp_ineq: "int p * k \<le> (int p * int q - 1) div 2 \<longleftrightarrow> k \<le> (int q - 1) div 2" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 52 | proof - | 
| 65413 | 53 | have "2 * int p * k \<le> int p * int q - 1 \<longleftrightarrow> 2 * k \<le> int q - 1" | 
| 54 | using p_ge_0 by auto | |
| 55 | then show ?thesis by auto | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 56 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 57 | |
| 65413 | 58 | lemma QRqp: "QR q p" | 
| 59 | using QR_def QR_axioms by simp | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 60 | |
| 65413 | 61 | lemma pq_commute: "int p * int q = int q * int p" | 
| 62 | by simp | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 63 | |
| 65413 | 64 | lemma pq_ge_0: "int p * int q > 0" | 
| 65 | using p_ge_0 q_ge_0 mult_pos_pos by blast | |
| 66 | ||
| 67 | definition "r = ((p - 1) div 2) * ((q - 1) div 2)" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 68 | definition "m = card (GAUSS.E p q)" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 69 | definition "n = card (GAUSS.E q p)" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 70 | |
| 65413 | 71 | abbreviation "Res k \<equiv> {0 .. k - 1}" for k :: int
 | 
| 72 | abbreviation "Res_ge_0 k \<equiv> {0 <.. k - 1}" for k :: int
 | |
| 73 | abbreviation "Res_0 k \<equiv> {0::int}" for k :: int
 | |
| 74 | abbreviation "Res_l k \<equiv> {0 <.. (k - 1) div 2}" for k :: int
 | |
| 75 | abbreviation "Res_h k \<equiv> {(k - 1) div 2 <.. k - 1}" for k :: int
 | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 76 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 77 | abbreviation "Sets_pq r0 r1 r2 \<equiv> | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 78 |   {(x::int). x \<in> r0 (int p * int q) \<and> x mod p \<in> r1 (int p) \<and> x mod q \<in> r2 (int q)}"
 | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 79 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 80 | definition "A = Sets_pq Res_l Res_l Res_h" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 81 | definition "B = Sets_pq Res_l Res_h Res_l" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 82 | definition "C = Sets_pq Res_h Res_h Res_l" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 83 | definition "D = Sets_pq Res_l Res_h Res_h" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 84 | definition "E = Sets_pq Res_l Res_0 Res_h" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 85 | definition "F = Sets_pq Res_l Res_h Res_0" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 86 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 87 | definition "a = card A" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 88 | definition "b = card B" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 89 | definition "c = card C" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 90 | definition "d = card D" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 91 | definition "e = card E" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 92 | definition "f = card F" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 93 | |
| 65413 | 94 | lemma Gpq: "GAUSS p q" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 95 | using p_prime pq_neq p_ge_2 q_prime | 
| 68707 
5b269897df9d
de-applying and removal of obsolete aliases
 paulson <lp15@cam.ac.uk> parents: 
68634diff
changeset | 96 | by (auto simp: GAUSS_def cong_iff_dvd_diff dest: primes_dvd_imp_eq) | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 97 | |
| 65413 | 98 | lemma Gqp: "GAUSS q p" | 
| 99 | by (simp add: QRqp QR.Gpq) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 100 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 101 | lemma QR_lemma_01: "(\<lambda>x. x mod q) ` E = GAUSS.E q p" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 102 | proof | 
| 65413 | 103 | have "x \<in> E \<longrightarrow> x mod int q \<in> GAUSS.E q p" if "x \<in> E" for x | 
| 104 | proof - | |
| 105 | from that obtain k where k: "x = int p * k" | |
| 106 | unfolding E_def by blast | |
| 107 | from that E_def have "x \<in> Res_l (int p * int q)" | |
| 108 | by blast | |
| 109 | then have "k \<in> GAUSS.A q" | |
| 110 | using Gqp GAUSS.A_def k qp_ineq by (simp add: zero_less_mult_iff) | |
| 111 | then have "x mod q \<in> GAUSS.E q p" | |
| 112 | using GAUSS.C_def[of q p] Gqp k GAUSS.B_def[of q p] that GAUSS.E_def[of q p] | |
| 113 | by (force simp: E_def) | |
| 114 | then show ?thesis by auto | |
| 115 | qed | |
| 116 | then show "(\<lambda>x. x mod int q) ` E \<subseteq> GAUSS.E q p" | |
| 117 | by auto | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 118 | show "GAUSS.E q p \<subseteq> (\<lambda>x. x mod q) ` E" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 119 | proof | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 120 | fix x | 
| 65413 | 121 | assume x: "x \<in> GAUSS.E q p" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 122 | then obtain ka where ka: "ka \<in> GAUSS.A q" "x = (ka * p) mod q" | 
| 65413 | 123 | by (auto simp: Gqp GAUSS.B_def GAUSS.C_def GAUSS.E_def) | 
| 124 | then have "ka * p \<in> Res_l (int p * int q)" | |
| 125 | using Gqp p_ge_0 qp_ineq by (simp add: GAUSS.A_def Groups.mult_ac(2)) | |
| 126 | then show "x \<in> (\<lambda>x. x mod q) ` E" | |
| 127 | using ka x Gqp q_ge_0 by (force simp: E_def GAUSS.E_def) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 128 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 129 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 130 | |
| 65413 | 131 | lemma QR_lemma_02: "e = n" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 132 | proof - | 
| 65413 | 133 | have "x = y" if x: "x \<in> E" and y: "y \<in> E" and mod: "x mod q = y mod q" for x y | 
| 134 | proof - | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 135 | obtain p_inv where p_inv: "[int p * p_inv = 1] (mod int q)" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 136 | using pq_coprime_int cong_solve_coprime_int by blast | 
| 65413 | 137 | from x y E_def obtain kx ky where k: "x = int p * kx" "y = int p * ky" | 
| 138 | using dvd_def[of p x] by blast | |
| 139 | with x y E_def have "0 < x" "int p * kx \<le> (int p * int q - 1) div 2" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 140 | "0 < y" "int p * ky \<le> (int p * int q - 1) div 2" | 
| 65413 | 141 | using greaterThanAtMost_iff mem_Collect_eq by blast+ | 
| 142 | with k have "0 \<le> kx" "kx < q" "0 \<le> ky" "ky < q" | |
| 143 | using qp_ineq by (simp_all add: zero_less_mult_iff) | |
| 144 | moreover from mod k have "(p_inv * (p * kx)) mod q = (p_inv * (p * ky)) mod q" | |
| 145 | using mod_mult_cong by blast | |
| 146 | then have "(p * p_inv * kx) mod q = (p * p_inv * ky) mod q" | |
| 147 | by (simp add: algebra_simps) | |
| 148 | then have "kx mod q = ky mod q" | |
| 66888 | 149 | using p_inv mod_mult_cong[of "p * p_inv" "q" "1"] | 
| 150 | by (auto simp: cong_def) | |
| 65413 | 151 | then have "[kx = ky] (mod q)" | 
| 66888 | 152 | unfolding cong_def by blast | 
| 65413 | 153 | ultimately show ?thesis | 
| 154 | using cong_less_imp_eq_int k by blast | |
| 155 | qed | |
| 156 | then have "inj_on (\<lambda>x. x mod q) E" | |
| 157 | by (auto simp: inj_on_def) | |
| 158 | then show ?thesis | |
| 159 | using QR_lemma_01 card_image e_def n_def by fastforce | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 160 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 161 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 162 | lemma QR_lemma_03: "f = m" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 163 | proof - | 
| 65413 | 164 | have "F = QR.E q p" | 
| 165 | unfolding F_def pq_commute using QRqp QR.E_def[of q p] by fastforce | |
| 166 | then have "f = QR.e q p" | |
| 167 | unfolding f_def using QRqp QR.e_def[of q p] by presburger | |
| 168 | then show ?thesis | |
| 169 | using QRqp QR.QR_lemma_02 m_def QRqp QR.n_def by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 170 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 171 | |
| 65413 | 172 | definition f_1 :: "int \<Rightarrow> int \<times> int" | 
| 173 | where "f_1 x = ((x mod p), (x mod q))" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 174 | |
| 65413 | 175 | definition P_1 :: "int \<times> int \<Rightarrow> int \<Rightarrow> bool" | 
| 176 | where "P_1 res x \<longleftrightarrow> x mod p = fst res \<and> x mod q = snd res \<and> x \<in> Res (int p * int q)" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 177 | |
| 65413 | 178 | definition g_1 :: "int \<times> int \<Rightarrow> int" | 
| 179 | where "g_1 res = (THE x. P_1 res x)" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 180 | |
| 65413 | 181 | lemma P_1_lemma: | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 182 | fixes res :: "int \<times> int" | 
| 65413 | 183 | assumes "0 \<le> fst res" "fst res < p" "0 \<le> snd res" "snd res < q" | 
| 184 | shows "\<exists>!x. P_1 res x" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 185 | proof - | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 186 | obtain y k1 k2 where yk: "y = nat (fst res) + k1 * p" "y = nat (snd res) + k2 * q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 187 | using chinese_remainder[of p q] pq_coprime p_ge_0 q_ge_0 by fastforce | 
| 66888 | 188 | have "fst res = int (y - k1 * p)" | 
| 189 | using \<open>0 \<le> fst res\<close> yk(1) by simp | |
| 190 | moreover have "snd res = int (y - k2 * q)" | |
| 191 | using \<open>0 \<le> snd res\<close> yk(2) by simp | |
| 192 | ultimately have res: "res = (int (y - k1 * p), int (y - k2 * q))" | |
| 193 | by (simp add: prod_eq_iff) | |
| 194 | have y: "k1 * p \<le> y" "k2 * q \<le> y" | |
| 195 | using yk by simp_all | |
| 196 | from y have *: "[y = nat (fst res)] (mod p)" "[y = nat (snd res)] (mod q)" | |
| 197 | by (auto simp add: res cong_le_nat intro: exI [of _ k1] exI [of _ k2]) | |
| 198 | from * have "(y mod (int p * int q)) mod int p = fst res" "(y mod (int p * int q)) mod int q = snd res" | |
| 199 | using y apply (auto simp add: res of_nat_mult [symmetric] of_nat_mod [symmetric] mod_mod_cancel simp del: of_nat_mult) | |
| 200 | apply (metis \<open>fst res = int (y - k1 * p)\<close> assms(1) assms(2) cong_def mod_pos_pos_trivial nat_int of_nat_mod) | |
| 201 | apply (metis \<open>snd res = int (y - k2 * q)\<close> assms(3) assms(4) cong_def mod_pos_pos_trivial nat_int of_nat_mod) | |
| 202 | done | |
| 65413 | 203 | then obtain x where "P_1 res x" | 
| 204 | unfolding P_1_def | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 205 | using Divides.pos_mod_bound Divides.pos_mod_sign pq_ge_0 by fastforce | 
| 65413 | 206 | moreover have "a = b" if "P_1 res a" "P_1 res b" for a b | 
| 207 | proof - | |
| 208 | from that have "int p * int q dvd a - b" | |
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64318diff
changeset | 209 | using divides_mult[of "int p" "a - b" "int q"] pq_coprime_int mod_eq_dvd_iff [of a _ b] | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 210 | unfolding P_1_def by force | 
| 65413 | 211 | with that show ?thesis | 
| 212 | using dvd_imp_le_int[of "a - b"] unfolding P_1_def by fastforce | |
| 213 | qed | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 214 | ultimately show ?thesis by auto | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 215 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 216 | |
| 65413 | 217 | lemma g_1_lemma: | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 218 | fixes res :: "int \<times> int" | 
| 65413 | 219 | assumes "0 \<le> fst res" "fst res < p" "0 \<le> snd res" "snd res < q" | 
| 220 | shows "P_1 res (g_1 res)" | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 221 | using assms P_1_lemma [of res] theI' [of "P_1 res"] g_1_def | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 222 | by auto | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 223 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 224 | definition "BuC = Sets_pq Res_ge_0 Res_h Res_l" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 225 | |
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 226 | lemma finite_BuC [simp]: | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 227 | "finite BuC" | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 228 | proof - | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 229 |   {
 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 230 | fix p q :: nat | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 231 |     have "finite {x. 0 < x \<and> x < int p * int q}"
 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 232 | by simp | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 233 |     then have "finite {x.
 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 234 | 0 < x \<and> | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 235 | x < int p * int q \<and> | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 236 | (int p - 1) div 2 | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 237 | < x mod int p \<and> | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 238 | x mod int p < int p \<and> | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 239 | 0 < x mod int q \<and> | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 240 | x mod int q \<le> (int q - 1) div 2}" | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 241 | by (auto intro: rev_finite_subset) | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 242 | } | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 243 | then show ?thesis | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 244 | by (simp add: BuC_def) | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 245 | qed | 
| 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 246 | |
| 65413 | 247 | lemma QR_lemma_04: "card BuC = card (Res_h p \<times> Res_l q)" | 
| 248 | using card_bij_eq[of f_1 "BuC" "Res_h p \<times> Res_l q" g_1] | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 249 | proof | 
| 65413 | 250 | show "inj_on f_1 BuC" | 
| 251 | proof | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 252 | fix x y | 
| 65413 | 253 | assume *: "x \<in> BuC" "y \<in> BuC" "f_1 x = f_1 y" | 
| 254 | then have "int p * int q dvd x - y" | |
| 255 | using f_1_def pq_coprime_int divides_mult[of "int p" "x - y" "int q"] | |
| 256 | mod_eq_dvd_iff[of x _ y] | |
| 257 | by auto | |
| 258 | with * show "x = y" | |
| 259 | using dvd_imp_le_int[of "x - y" "int p * int q"] unfolding BuC_def by force | |
| 260 | qed | |
| 261 | show "inj_on g_1 (Res_h p \<times> Res_l q)" | |
| 262 | proof | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 263 | fix x y | 
| 65413 | 264 | assume *: "x \<in> Res_h p \<times> Res_l q" "y \<in> Res_h p \<times> Res_l q" "g_1 x = g_1 y" | 
| 265 | then have "0 \<le> fst x" "fst x < p" "0 \<le> snd x" "snd x < q" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 266 | "0 \<le> fst y" "fst y < p" "0 \<le> snd y" "snd y < q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 267 | using mem_Sigma_iff prod.collapse by fastforce+ | 
| 65413 | 268 | with * show "x = y" | 
| 269 | using g_1_lemma[of x] g_1_lemma[of y] P_1_def by fastforce | |
| 270 | qed | |
| 271 | show "g_1 ` (Res_h p \<times> Res_l q) \<subseteq> BuC" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 272 | proof | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 273 | fix y | 
| 65413 | 274 | assume "y \<in> g_1 ` (Res_h p \<times> Res_l q)" | 
| 275 | then obtain x where x: "y = g_1 x" "x \<in> Res_h p \<times> Res_l q" | |
| 276 | by blast | |
| 277 | then have "P_1 x y" | |
| 278 | using g_1_lemma by fastforce | |
| 279 | with x show "y \<in> BuC" | |
| 280 | unfolding P_1_def BuC_def mem_Collect_eq using SigmaE prod.sel by fastforce | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 281 | qed | 
| 65416 
f707dbcf11e3
more approproiate placement of theories MiscAlgebra and Multiplicate_Group
 haftmann parents: 
65413diff
changeset | 282 | qed (auto simp: finite_subset f_1_def, simp_all add: BuC_def) | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 283 | |
| 65413 | 284 | lemma QR_lemma_05: "card (Res_h p \<times> Res_l q) = r" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 285 | proof - | 
| 65413 | 286 | have "card (Res_l q) = (q - 1) div 2" "card (Res_h p) = (p - 1) div 2" | 
| 287 | using p_eq2 by force+ | |
| 288 | then show ?thesis | |
| 289 | unfolding r_def using card_cartesian_product[of "Res_h p" "Res_l q"] by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 290 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 291 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 292 | lemma QR_lemma_06: "b + c = r" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 293 | proof - | 
| 65413 | 294 |   have "B \<inter> C = {}" "finite B" "finite C" "B \<union> C = BuC"
 | 
| 295 | unfolding B_def C_def BuC_def by fastforce+ | |
| 296 | then show ?thesis | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 297 | unfolding b_def c_def using card_empty card_Un_Int QR_lemma_04 QR_lemma_05 by fastforce | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 298 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 299 | |
| 65413 | 300 | definition f_2:: "int \<Rightarrow> int" | 
| 301 | where "f_2 x = (int p * int q) - x" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 302 | |
| 65413 | 303 | lemma f_2_lemma_1: "f_2 (f_2 x) = x" | 
| 304 | by (simp add: f_2_def) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 305 | |
| 65413 | 306 | lemma f_2_lemma_2: "[f_2 x = int p - x] (mod p)" | 
| 68707 
5b269897df9d
de-applying and removal of obsolete aliases
 paulson <lp15@cam.ac.uk> parents: 
68634diff
changeset | 307 | by (simp add: f_2_def cong_iff_dvd_diff) | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 308 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 309 | lemma f_2_lemma_3: "f_2 x \<in> S \<Longrightarrow> x \<in> f_2 ` S" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 310 | using f_2_lemma_1[of x] image_eqI[of x f_2 "f_2 x" S] by presburger | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 311 | |
| 65413 | 312 | lemma QR_lemma_07: | 
| 313 | "f_2 ` Res_l (int p * int q) = Res_h (int p * int q)" | |
| 314 | "f_2 ` Res_h (int p * int q) = Res_l (int p * int q)" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 315 | proof - | 
| 65413 | 316 | have 1: "f_2 ` Res_l (int p * int q) \<subseteq> Res_h (int p * int q)" | 
| 317 | by (force simp: f_2_def) | |
| 318 | have 2: "f_2 ` Res_h (int p * int q) \<subseteq> Res_l (int p * int q)" | |
| 319 | using pq_eq2 by (fastforce simp: f_2_def) | |
| 320 | from 2 have 3: "Res_h (int p * int q) \<subseteq> f_2 ` Res_l (int p * int q)" | |
| 321 | using f_2_lemma_3 by blast | |
| 322 | from 1 have 4: "Res_l (int p * int q) \<subseteq> f_2 ` Res_h (int p * int q)" | |
| 323 | using f_2_lemma_3 by blast | |
| 324 | from 1 3 show "f_2 ` Res_l (int p * int q) = Res_h (int p * int q)" | |
| 325 | by blast | |
| 326 | from 2 4 show "f_2 ` Res_h (int p * int q) = Res_l (int p * int q)" | |
| 327 | by blast | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 328 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 329 | |
| 65413 | 330 | lemma QR_lemma_08: | 
| 331 | "f_2 x mod p \<in> Res_l p \<longleftrightarrow> x mod p \<in> Res_h p" | |
| 332 | "f_2 x mod p \<in> Res_h p \<longleftrightarrow> x mod p \<in> Res_l p" | |
| 66888 | 333 | using f_2_lemma_2[of x] cong_def[of "f_2 x" "p - x" p] minus_mod_self2[of x p] | 
| 65413 | 334 | zmod_zminus1_eq_if[of x p] p_eq2 | 
| 335 | by auto | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 336 | |
| 65413 | 337 | lemma QR_lemma_09: | 
| 338 | "f_2 x mod q \<in> Res_l q \<longleftrightarrow> x mod q \<in> Res_h q" | |
| 339 | "f_2 x mod q \<in> Res_h q \<longleftrightarrow> x mod q \<in> Res_l q" | |
| 340 | using QRqp QR.QR_lemma_08 f_2_def QR.f_2_def pq_commute by auto | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 341 | |
| 65413 | 342 | lemma QR_lemma_10: "a = c" | 
| 343 | unfolding a_def c_def | |
| 344 | apply (rule card_bij_eq[of f_2 A C f_2]) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 345 | unfolding A_def C_def | 
| 65413 | 346 | using QR_lemma_07 QR_lemma_08 QR_lemma_09 apply ((simp add: inj_on_def f_2_def), blast)+ | 
| 347 | apply fastforce+ | |
| 348 | done | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 349 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 350 | definition "BuD = Sets_pq Res_l Res_h Res_ge_0" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 351 | definition "BuDuF = Sets_pq Res_l Res_h Res" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 352 | |
| 65413 | 353 | definition f_3 :: "int \<Rightarrow> int \<times> int" | 
| 354 | where "f_3 x = (x mod p, x div p + 1)" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 355 | |
| 65413 | 356 | definition g_3 :: "int \<times> int \<Rightarrow> int" | 
| 357 | where "g_3 x = fst x + (snd x - 1) * p" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 358 | |
| 65413 | 359 | lemma QR_lemma_11: "card BuDuF = card (Res_h p \<times> Res_l q)" | 
| 360 | using card_bij_eq[of f_3 BuDuF "Res_h p \<times> Res_l q" g_3] | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 361 | proof | 
| 65413 | 362 | show "f_3 ` BuDuF \<subseteq> Res_h p \<times> Res_l q" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 363 | proof | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 364 | fix y | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 365 | assume "y \<in> f_3 ` BuDuF" | 
| 65413 | 366 | then obtain x where x: "y = f_3 x" "x \<in> BuDuF" | 
| 367 | by blast | |
| 368 | then have "x \<le> int p * (int q - 1) div 2 + (int p - 1) div 2" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 369 | unfolding BuDuF_def using p_eq2 int_distrib(4) by auto | 
| 65413 | 370 | moreover from x have "(int p - 1) div 2 \<le> - 1 + x mod p" | 
| 371 | by (auto simp: BuDuF_def) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 372 | moreover have "int p * (int q - 1) div 2 = int p * ((int q - 1) div 2)" | 
| 68634 | 373 | by (subst div_mult1_eq) (simp add: odd_q) | 
| 65413 | 374 | then have "p * (int q - 1) div 2 = p * ((int q + 1) div 2 - 1)" | 
| 375 | by fastforce | |
| 376 | ultimately have "x \<le> p * ((int q + 1) div 2 - 1) - 1 + x mod p" | |
| 377 | by linarith | |
| 378 | then have "x div p < (int q + 1) div 2 - 1" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 379 | using mult.commute[of "int p" "x div p"] p_ge_0 div_mult_mod_eq[of x p] | 
| 65413 | 380 | and mult_less_cancel_left_pos[of p "x div p" "(int q + 1) div 2 - 1"] | 
| 381 | by linarith | |
| 382 | moreover from x have "0 < x div p + 1" | |
| 383 | using pos_imp_zdiv_neg_iff[of p x] p_ge_0 by (auto simp: BuDuF_def) | |
| 384 | ultimately show "y \<in> Res_h p \<times> Res_l q" | |
| 385 | using x by (auto simp: BuDuF_def f_3_def) | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 386 | qed | 
| 65413 | 387 | show "inj_on g_3 (Res_h p \<times> Res_l q)" | 
| 388 | proof | |
| 389 | have *: "f_3 (g_3 x) = x" if "x \<in> Res_h p \<times> Res_l q" for x | |
| 390 | proof - | |
| 391 | from that have *: "(fst x + (snd x - 1) * int p) mod int p = fst x" | |
| 392 | by force | |
| 393 | from that have "(fst x + (snd x - 1) * int p) div int p + 1 = snd x" | |
| 66888 | 394 | by auto | 
| 65413 | 395 | with * show "f_3 (g_3 x) = x" | 
| 396 | by (simp add: f_3_def g_3_def) | |
| 397 | qed | |
| 398 | fix x y | |
| 399 | assume "x \<in> Res_h p \<times> Res_l q" "y \<in> Res_h p \<times> Res_l q" "g_3 x = g_3 y" | |
| 400 | from this *[of x] *[of y] show "x = y" | |
| 401 | by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 402 | qed | 
| 65413 | 403 | show "g_3 ` (Res_h p \<times> Res_l q) \<subseteq> BuDuF" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 404 | proof | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 405 | fix y | 
| 65413 | 406 | assume "y \<in> g_3 ` (Res_h p \<times> Res_l q)" | 
| 407 | then obtain x where x: "x \<in> Res_h p \<times> Res_l q" and y: "y = g_3 x" | |
| 408 | by blast | |
| 409 | then have "snd x \<le> (int q - 1) div 2" | |
| 410 | by force | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 411 | moreover have "int p * ((int q - 1) div 2) = (int p * int q - int p) div 2" | 
| 68631 | 412 | using int_distrib(4) div_mult1_eq[of "int p" "int q - 1" 2] odd_q by fastforce | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 413 | ultimately have "(snd x) * int p \<le> (int q * int p - int p) div 2" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 414 | using mult_right_mono[of "snd x" "(int q - 1) div 2" p] mult.commute[of "(int q - 1) div 2" p] | 
| 65413 | 415 | pq_commute | 
| 416 | by presburger | |
| 417 | then have "(snd x - 1) * int p \<le> (int q * int p - 1) div 2 - int p" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 418 | using p_ge_0 int_distrib(3) by auto | 
| 65413 | 419 | moreover from x have "fst x \<le> int p - 1" by force | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 420 | ultimately have "fst x + (snd x - 1) * int p \<le> (int p * int q - 1) div 2" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 421 | using pq_commute by linarith | 
| 65413 | 422 | moreover from x have "0 < fst x" "0 \<le> (snd x - 1) * p" | 
| 423 | by fastforce+ | |
| 424 | ultimately show "y \<in> BuDuF" | |
| 425 | unfolding BuDuF_def using q_ge_0 x g_3_def y by auto | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 426 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 427 | show "finite BuDuF" unfolding BuDuF_def by fastforce | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 428 | qed (simp add: inj_on_inverseI[of BuDuF g_3] f_3_def g_3_def QR_lemma_05)+ | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 429 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 430 | lemma QR_lemma_12: "b + d + m = r" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 431 | proof - | 
| 65413 | 432 |   have "B \<inter> D = {}" "finite B" "finite D" "B \<union> D = BuD"
 | 
| 433 | unfolding B_def D_def BuD_def by fastforce+ | |
| 434 | then have "b + d = card BuD" | |
| 435 | unfolding b_def d_def using card_Un_Int by fastforce | |
| 436 |   moreover have "BuD \<inter> F = {}" "finite BuD" "finite F"
 | |
| 437 | unfolding BuD_def F_def by fastforce+ | |
| 438 | moreover have "BuD \<union> F = BuDuF" | |
| 439 | unfolding BuD_def F_def BuDuF_def | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 440 | using q_ge_0 ivl_disj_un_singleton(5)[of 0 "int q - 1"] by auto | 
| 65413 | 441 | ultimately show ?thesis | 
| 442 | using QR_lemma_03 QR_lemma_05 QR_lemma_11 card_Un_disjoint[of BuD F] | |
| 443 | unfolding b_def d_def f_def | |
| 444 | by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 445 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 446 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 447 | lemma QR_lemma_13: "a + d + n = r" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 448 | proof - | 
| 65413 | 449 | have "A = QR.B q p" | 
| 450 | unfolding A_def pq_commute using QRqp QR.B_def[of q p] by blast | |
| 451 | then have "a = QR.b q p" | |
| 452 | using a_def QRqp QR.b_def[of q p] by presburger | |
| 453 | moreover have "D = QR.D q p" | |
| 454 | unfolding D_def pq_commute using QRqp QR.D_def[of q p] by blast | |
| 455 | then have "d = QR.d q p" | |
| 456 | using d_def QRqp QR.d_def[of q p] by presburger | |
| 457 | moreover have "n = QR.m q p" | |
| 458 | using n_def QRqp QR.m_def[of q p] by presburger | |
| 459 | moreover have "r = QR.r q p" | |
| 460 | unfolding r_def using QRqp QR.r_def[of q p] by auto | |
| 461 | ultimately show ?thesis | |
| 462 | using QRqp QR.QR_lemma_12 by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 463 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 464 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 465 | lemma QR_lemma_14: "(-1::int) ^ (m + n) = (-1) ^ r" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 466 | proof - | 
| 65413 | 467 | have "m + n + 2 * d = r" | 
| 468 | using QR_lemma_06 QR_lemma_10 QR_lemma_12 QR_lemma_13 by auto | |
| 469 | then show ?thesis | |
| 470 | using power_add[of "-1::int" "m + n" "2 * d"] by fastforce | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 471 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 472 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 473 | lemma Quadratic_Reciprocity: | 
| 65413 | 474 | "Legendre p q * Legendre q p = (-1::int) ^ ((p - 1) div 2 * ((q - 1) div 2))" | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 475 | using Gpq Gqp GAUSS.gauss_lemma power_add[of "-1::int" m n] QR_lemma_14 | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 476 | unfolding r_def m_def n_def by auto | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 477 | |
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 478 | end | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 479 | |
| 65413 | 480 | theorem Quadratic_Reciprocity: | 
| 481 | assumes "prime p" "2 < p" "prime q" "2 < q" "p \<noteq> q" | |
| 482 | shows "Legendre p q * Legendre q p = (-1::int) ^ ((p - 1) div 2 * ((q - 1) div 2))" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 483 | using QR.Quadratic_Reciprocity QR_def assms by blast | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 484 | |
| 65413 | 485 | theorem Quadratic_Reciprocity_int: | 
| 486 | assumes "prime (nat p)" "2 < p" "prime (nat q)" "2 < q" "p \<noteq> q" | |
| 487 | shows "Legendre p q * Legendre q p = (-1::int) ^ (nat ((p - 1) div 2 * ((q - 1) div 2)))" | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 488 | proof - | 
| 65413 | 489 | from assms have "0 \<le> (p - 1) div 2" by simp | 
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 490 | moreover have "(nat p - 1) div 2 = nat ((p - 1) div 2)" "(nat q - 1) div 2 = nat ((q - 1) div 2)" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 491 | by fastforce+ | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 492 | ultimately have "(nat p - 1) div 2 * ((nat q - 1) div 2) = nat ((p - 1) div 2 * ((q - 1) div 2))" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 493 | using nat_mult_distrib by presburger | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 494 | moreover have "2 < nat p" "2 < nat q" "nat p \<noteq> nat q" "int (nat p) = p" "int (nat q) = q" | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 495 | using assms by linarith+ | 
| 65413 | 496 | ultimately show ?thesis | 
| 497 | using Quadratic_Reciprocity[of "nat p" "nat q"] assms by presburger | |
| 64282 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 498 | qed | 
| 
261d42f0bfac
Removed Old_Number_Theory; all theories ported (thanks to Jaime Mendizabal Roche)
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 499 | |
| 64318 
1e92b5c35615
Repaired LaTeX in HOL-Data_Structures
 eberlm <eberlm@in.tum.de> parents: 
64282diff
changeset | 500 | end |