| author | berghofe | 
| Mon, 23 Oct 2006 00:51:16 +0200 | |
| changeset 21087 | 3e56528a39f7 | 
| parent 20432 | 07ec57376051 | 
| child 21256 | 47195501ecf7 | 
| permissions | -rw-r--r-- | 
| 14706 | 1 | (* Title: HOL/Algebra/Exponent.thy | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 3 | Author: Florian Kammueller, with new proofs by L C Paulson | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 4 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 5 | exponent p s yields the greatest power of p that divides s. | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 6 | *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 7 | |
| 16417 | 8 | theory Exponent imports Main Primes begin | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 9 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20282diff
changeset | 10 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20282diff
changeset | 11 | section {*The Combinatorial Argument Underlying the First Sylow Theorem*}
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 12 | constdefs | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 13 | exponent :: "[nat, nat] => nat" | 
| 16663 | 14 | "exponent p s == if prime p then (GREATEST r. p^r dvd s) else 0" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 15 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20282diff
changeset | 16 | |
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 17 | subsection{*Prime Theorems*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 18 | |
| 16663 | 19 | lemma prime_imp_one_less: "prime p ==> Suc 0 < p" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 20 | by (unfold prime_def, force) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 21 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 22 | lemma prime_iff: | 
| 16663 | 23 | "(prime p) = (Suc 0 < p & (\<forall>a b. p dvd a*b --> (p dvd a) | (p dvd b)))" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 24 | apply (auto simp add: prime_imp_one_less) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 25 | apply (blast dest!: prime_dvd_mult) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 26 | apply (auto simp add: prime_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 27 | apply (erule dvdE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 28 | apply (case_tac "k=0", simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 29 | apply (drule_tac x = m in spec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 30 | apply (drule_tac x = k in spec) | 
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
16663diff
changeset | 31 | apply (simp add: dvd_mult_cancel1 dvd_mult_cancel2) | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 32 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 33 | |
| 16663 | 34 | lemma zero_less_prime_power: "prime p ==> 0 < p^a" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 35 | by (force simp add: prime_iff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 36 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 37 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 38 | lemma zero_less_card_empty: "[| finite S; S \<noteq> {} |] ==> 0 < card(S)"
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 39 | by (rule ccontr, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 40 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 41 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 42 | lemma prime_dvd_cases: | 
| 16663 | 43 | "[| p*k dvd m*n; prime p |] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 44 | ==> (\<exists>x. k dvd x*n & m = p*x) | (\<exists>y. k dvd m*y & n = p*y)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 45 | apply (simp add: prime_iff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 46 | apply (frule dvd_mult_left) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 47 | apply (subgoal_tac "p dvd m | p dvd n") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 48 | prefer 2 apply blast | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 49 | apply (erule disjE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 50 | apply (rule disjI1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 51 | apply (rule_tac [2] disjI2) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 52 | apply (erule_tac n = m in dvdE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 53 | apply (erule_tac [2] n = n in dvdE, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 54 | apply (rule_tac [2] k = p in dvd_mult_cancel) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 55 | apply (rule_tac k = p in dvd_mult_cancel) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 56 | apply (simp_all add: mult_ac) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 57 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 58 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 59 | |
| 16663 | 60 | lemma prime_power_dvd_cases [rule_format (no_asm)]: "prime p | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 61 | ==> \<forall>m n. p^c dvd m*n --> | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 62 | (\<forall>a b. a+b = Suc c --> p^a dvd m | p^b dvd n)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 63 | apply (induct_tac "c") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 64 | apply clarify | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 65 | apply (case_tac "a") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 66 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 67 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 68 | (*inductive step*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 69 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 70 | apply clarify | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 71 | apply (erule prime_dvd_cases [THEN disjE], assumption, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 72 | (*case 1: p dvd m*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 73 | apply (case_tac "a") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 74 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 75 | apply clarify | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 76 | apply (drule spec, drule spec, erule (1) notE impE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 77 | apply (drule_tac x = nat in spec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 78 | apply (drule_tac x = b in spec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 79 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 80 | apply (blast intro: dvd_refl mult_dvd_mono) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 81 | (*case 2: p dvd n*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 82 | apply (case_tac "b") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 83 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 84 | apply clarify | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 85 | apply (drule spec, drule spec, erule (1) notE impE) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 86 | apply (drule_tac x = a in spec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 87 | apply (drule_tac x = nat in spec, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 88 | apply (blast intro: dvd_refl mult_dvd_mono) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 89 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 90 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 91 | (*needed in this form in Sylow.ML*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 92 | lemma div_combine: | 
| 16663 | 93 | "[| prime p; ~ (p ^ (Suc r) dvd n); p^(a+r) dvd n*k |] | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 94 | ==> p ^ a dvd k" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 95 | by (drule_tac a = "Suc r" and b = a in prime_power_dvd_cases, assumption, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 96 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 97 | (*Lemma for power_dvd_bound*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 98 | lemma Suc_le_power: "Suc 0 < p ==> Suc n <= p^n" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 99 | apply (induct_tac "n") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 100 | apply (simp (no_asm_simp)) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 101 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 102 | apply (subgoal_tac "2 * n + 2 <= p * p^n", simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 103 | apply (subgoal_tac "2 * p^n <= p * p^n") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 104 | (*?arith_tac should handle all of this!*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 105 | apply (rule order_trans) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 106 | prefer 2 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 107 | apply (drule_tac k = 2 in mult_le_mono2, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 108 | apply (rule mult_le_mono1, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 109 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 110 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 111 | (*An upper bound for the n such that p^n dvd a: needed for GREATEST to exist*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 112 | lemma power_dvd_bound: "[|p^n dvd a; Suc 0 < p; 0 < a|] ==> n < a" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 113 | apply (drule dvd_imp_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 114 | apply (drule_tac [2] n = n in Suc_le_power, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 115 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 116 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 117 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 118 | subsection{*Exponent Theorems*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 119 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 120 | lemma exponent_ge [rule_format]: | 
| 16663 | 121 | "[|p^k dvd n; prime p; 0<n|] ==> k <= exponent p n" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 122 | apply (simp add: exponent_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 123 | apply (erule Greatest_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 124 | apply (blast dest: prime_imp_one_less power_dvd_bound) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 125 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 126 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 127 | lemma power_exponent_dvd: "0<s ==> (p ^ exponent p s) dvd s" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 128 | apply (simp add: exponent_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 129 | apply clarify | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 130 | apply (rule_tac k = 0 in GreatestI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 131 | prefer 2 apply (blast dest: prime_imp_one_less power_dvd_bound, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 132 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 133 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 134 | lemma power_Suc_exponent_Not_dvd: | 
| 16663 | 135 | "[|(p * p ^ exponent p s) dvd s; prime p |] ==> s=0" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 136 | apply (subgoal_tac "p ^ Suc (exponent p s) dvd s") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 137 | prefer 2 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 138 | apply (rule ccontr) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 139 | apply (drule exponent_ge, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 140 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 141 | |
| 16663 | 142 | lemma exponent_power_eq [simp]: "prime p ==> exponent p (p^a) = a" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 143 | apply (simp (no_asm_simp) add: exponent_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 144 | apply (rule Greatest_equality, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 145 | apply (simp (no_asm_simp) add: prime_imp_one_less power_dvd_imp_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 146 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 147 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 148 | lemma exponent_equalityI: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 149 | "!r::nat. (p^r dvd a) = (p^r dvd b) ==> exponent p a = exponent p b" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 150 | by (simp (no_asm_simp) add: exponent_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 151 | |
| 16663 | 152 | lemma exponent_eq_0 [simp]: "\<not> prime p ==> exponent p s = 0" | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 153 | by (simp (no_asm_simp) add: exponent_def) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 154 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 155 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 156 | (* exponent_mult_add, easy inclusion. Could weaken p \<in> prime to Suc 0 < p *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 157 | lemma exponent_mult_add1: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 158 | "[| 0 < a; 0 < b |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 159 | ==> (exponent p a) + (exponent p b) <= exponent p (a * b)" | 
| 16663 | 160 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 161 | apply (rule exponent_ge) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 162 | apply (auto simp add: power_add) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 163 | apply (blast intro: prime_imp_one_less power_exponent_dvd mult_dvd_mono) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 164 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 165 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 166 | (* exponent_mult_add, opposite inclusion *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 167 | lemma exponent_mult_add2: "[| 0 < a; 0 < b |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 168 | ==> exponent p (a * b) <= (exponent p a) + (exponent p b)" | 
| 16663 | 169 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 170 | apply (rule leI, clarify) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 171 | apply (cut_tac p = p and s = "a*b" in power_exponent_dvd, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 172 | apply (subgoal_tac "p ^ (Suc (exponent p a + exponent p b)) dvd a * b") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 173 | apply (rule_tac [2] le_imp_power_dvd [THEN dvd_trans]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 174 | prefer 3 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 175 | prefer 2 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 176 | apply (frule_tac a = "Suc (exponent p a) " and b = "Suc (exponent p b) " in prime_power_dvd_cases) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 177 | apply (assumption, force, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 178 | apply (blast dest: power_Suc_exponent_Not_dvd) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 179 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 180 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 181 | lemma exponent_mult_add: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 182 | "[| 0 < a; 0 < b |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 183 | ==> exponent p (a * b) = (exponent p a) + (exponent p b)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 184 | by (blast intro: exponent_mult_add1 exponent_mult_add2 order_antisym) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 185 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 186 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 187 | lemma not_divides_exponent_0: "~ (p dvd n) ==> exponent p n = 0" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 188 | apply (case_tac "exponent p n", simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 189 | apply (case_tac "n", simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 190 | apply (cut_tac s = n and p = p in power_exponent_dvd) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 191 | apply (auto dest: dvd_mult_left) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 192 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 193 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 194 | lemma exponent_1_eq_0 [simp]: "exponent p (Suc 0) = 0" | 
| 16663 | 195 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 196 | apply (auto simp add: prime_iff not_divides_exponent_0) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 197 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 198 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 199 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: 
20282diff
changeset | 200 | subsection{*Main Combinatorial Argument*}
 | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 201 | |
| 14889 | 202 | lemma le_extend_mult: "[| 0 < c; a <= b |] ==> a <= b * (c::nat)" | 
| 203 | apply (rule_tac P = "%x. x <= b * c" in subst) | |
| 204 | apply (rule mult_1_right) | |
| 205 | apply (rule mult_le_mono, auto) | |
| 206 | done | |
| 207 | ||
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 208 | lemma p_fac_forw_lemma: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 209 | "[| 0 < (m::nat); 0<k; k < p^a; (p^r) dvd (p^a)* m - k |] ==> r <= a" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 210 | apply (rule notnotD) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 211 | apply (rule notI) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 212 | apply (drule contrapos_nn [OF _ leI, THEN notnotD], assumption) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 213 | apply (drule_tac m = a in less_imp_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 214 | apply (drule le_imp_power_dvd) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 215 | apply (drule_tac n = "p ^ r" in dvd_trans, assumption) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 216 | apply (frule_tac m = k in less_imp_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 217 | apply (drule_tac c = m in le_extend_mult, assumption) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 218 | apply (drule_tac k = "p ^ a" and m = " (p ^ a) * m" in dvd_diffD1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 219 | prefer 2 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 220 | apply (rule dvd_refl [THEN dvd_mult2]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 221 | apply (drule_tac n = k in dvd_imp_le, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 222 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 223 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 224 | lemma p_fac_forw: "[| 0 < (m::nat); 0<k; k < p^a; (p^r) dvd (p^a)* m - k |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 225 | ==> (p^r) dvd (p^a) - k" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 226 | apply (frule_tac k1 = k and i = p in p_fac_forw_lemma [THEN le_imp_power_dvd], auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 227 | apply (subgoal_tac "p^r dvd p^a*m") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 228 | prefer 2 apply (blast intro: dvd_mult2) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 229 | apply (drule dvd_diffD1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 230 | apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 231 | prefer 2 apply (blast intro: dvd_diff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 232 | apply (drule less_imp_Suc_add, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 233 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 234 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 235 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 236 | lemma r_le_a_forw: "[| 0 < (k::nat); k < p^a; 0 < p; (p^r) dvd (p^a) - k |] ==> r <= a" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 237 | by (rule_tac m = "Suc 0" in p_fac_forw_lemma, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 238 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 239 | lemma p_fac_backw: "[| 0<m; 0<k; 0 < (p::nat); k < p^a; (p^r) dvd p^a - k |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 240 | ==> (p^r) dvd (p^a)*m - k" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 241 | apply (frule_tac k1 = k and i = p in r_le_a_forw [THEN le_imp_power_dvd], auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 242 | apply (subgoal_tac "p^r dvd p^a*m") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 243 | prefer 2 apply (blast intro: dvd_mult2) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 244 | apply (drule dvd_diffD1) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 245 | apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 246 | prefer 2 apply (blast intro: dvd_diff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 247 | apply (drule less_imp_Suc_add, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 248 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 249 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 250 | lemma exponent_p_a_m_k_equation: "[| 0<m; 0<k; 0 < (p::nat); k < p^a |] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 251 | ==> exponent p (p^a * m - k) = exponent p (p^a - k)" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 252 | apply (blast intro: exponent_equalityI p_fac_forw p_fac_backw) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 253 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 254 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 255 | text{*Suc rules that we have to delete from the simpset*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 256 | lemmas bad_Sucs = binomial_Suc_Suc mult_Suc mult_Suc_right | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 257 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 258 | (*The bound K is needed; otherwise it's too weak to be used.*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 259 | lemma p_not_div_choose_lemma [rule_format]: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 260 | "[| \<forall>i. Suc i < K --> exponent p (Suc i) = exponent p (Suc(j+i))|] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 261 | ==> k<K --> exponent p ((j+k) choose k) = 0" | 
| 16663 | 262 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 263 | prefer 2 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 264 | apply (induct_tac "k") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 265 | apply (simp (no_asm)) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 266 | (*induction step*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 267 | apply (subgoal_tac "0 < (Suc (j+n) choose Suc n) ") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 268 | prefer 2 apply (simp add: zero_less_binomial_iff, clarify) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 269 | apply (subgoal_tac "exponent p ((Suc (j+n) choose Suc n) * Suc n) = | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 270 | exponent p (Suc n)") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 271 |  txt{*First, use the assumed equation.  We simplify the LHS to
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 272 |   @{term "exponent p (Suc (j + n) choose Suc n) + exponent p (Suc n)"}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 273 | the common terms cancel, proving the conclusion.*} | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 274 | apply (simp del: bad_Sucs add: exponent_mult_add) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 275 | txt{*Establishing the equation requires first applying 
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 276 |    @{text Suc_times_binomial_eq} ...*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 277 | apply (simp del: bad_Sucs add: Suc_times_binomial_eq [symmetric]) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 278 | txt{*...then @{text exponent_mult_add} and the quantified premise.*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 279 | apply (simp del: bad_Sucs add: zero_less_binomial_iff exponent_mult_add) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 280 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 281 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 282 | (*The lemma above, with two changes of variables*) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 283 | lemma p_not_div_choose: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 284 | "[| k<K; k<=n; | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 285 | \<forall>j. 0<j & j<K --> exponent p (n - k + (K - j)) = exponent p (K - j)|] | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 286 | ==> exponent p (n choose k) = 0" | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 287 | apply (cut_tac j = "n-k" and k = k and p = p in p_not_div_choose_lemma) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 288 | prefer 3 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 289 | prefer 2 apply assumption | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 290 | apply (drule_tac x = "K - Suc i" in spec) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 291 | apply (simp add: Suc_diff_le) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 292 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 293 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 294 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 295 | lemma const_p_fac_right: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 296 | "0 < m ==> exponent p ((p^a * m - Suc 0) choose (p^a - Suc 0)) = 0" | 
| 16663 | 297 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 298 | prefer 2 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 299 | apply (frule_tac a = a in zero_less_prime_power) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 300 | apply (rule_tac K = "p^a" in p_not_div_choose) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 301 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 302 | apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 303 | apply (case_tac "m") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 304 | apply (case_tac [2] "p^a") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 305 | apply auto | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 306 | (*now the hard case, simplified to | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 307 | exponent p (Suc (p ^ a * m + i - p ^ a)) = exponent p (Suc i) *) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 308 | apply (subgoal_tac "0<p") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 309 | prefer 2 apply (force dest!: prime_imp_one_less) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 310 | apply (subst exponent_p_a_m_k_equation, auto) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 311 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 312 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 313 | lemma const_p_fac: | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 314 | "0 < m ==> exponent p (((p^a) * m) choose p^a) = exponent p m" | 
| 16663 | 315 | apply (case_tac "prime p") | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 316 | prefer 2 apply simp | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 317 | apply (subgoal_tac "0 < p^a * m & p^a <= p^a * m") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 318 | prefer 2 apply (force simp add: prime_iff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 319 | txt{*A similar trick to the one used in @{text p_not_div_choose_lemma}:
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 320 |   insert an equation; use @{text exponent_mult_add} on the LHS; on the RHS,
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 321 | first | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 322 |   transform the binomial coefficient, then use @{text exponent_mult_add}.*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 323 | apply (subgoal_tac "exponent p ((( (p^a) * m) choose p^a) * p^a) = | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 324 | a + exponent p m") | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 325 | apply (simp del: bad_Sucs add: zero_less_binomial_iff exponent_mult_add prime_iff) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 326 | txt{*one subgoal left!*}
 | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 327 | apply (subst times_binomial_minus1_eq, simp, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 328 | apply (subst exponent_mult_add, simp) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 329 | apply (simp (no_asm_simp) add: zero_less_binomial_iff) | 
| 20432 
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
 webertj parents: 
20318diff
changeset | 330 | apply arith | 
| 13870 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 331 | apply (simp del: bad_Sucs add: exponent_mult_add const_p_fac_right) | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 332 | done | 
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 333 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 334 | |
| 
cf947d1ec5ff
moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
 paulson parents: diff
changeset | 335 | end |