src/HOL/Algebra/Exponent.thy
author paulson <lp15@cam.ac.uk>
Fri, 07 Nov 2014 15:40:08 +0000
changeset 58917 a3be9a47e2d7
parent 57865 dcfb33c26f50
child 59667 651ea265d568
permissions -rw-r--r--
Tidying up. Removing unnecessary conditions from some theorems.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14706
71590b7733b7 tuned document;
wenzelm
parents: 13870
diff changeset
     1
(*  Title:      HOL/Algebra/Exponent.thy
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     2
    Author:     Florian Kammueller
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     3
    Author:     L C Paulson
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
     4
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     5
exponent p s   yields the greatest power of p that divides s.
13870
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
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
     8
theory Exponent
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
     9
imports Main "~~/src/HOL/Number_Theory/Primes" "~~/src/HOL/Number_Theory/Binomial"
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
    10
begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
    11
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
    12
section {*Sylow's Theorem*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
    13
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
    14
subsection {*The Combinatorial Argument Underlying the First Sylow Theorem*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
    15
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 32946
diff changeset
    16
definition
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 32946
diff changeset
    17
  exponent :: "nat => nat => nat"
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 32946
diff changeset
    18
  where "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
    19
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
    20
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
    21
text{*Prime Theorems*}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    22
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    23
lemma prime_iff:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
    24
  "(prime p) = (Suc 0 < p & (\<forall>a b. p dvd a*b --> (p dvd a) | (p dvd b)))"
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    25
apply (auto simp add: prime_gt_Suc_0_nat)
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    26
by (metis (full_types) One_nat_def Suc_lessD dvd.order_refl nat_dvd_not_less not_prime_eq_prod_nat)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    27
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    28
lemma zero_less_prime_power:
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    29
  fixes p::nat shows "prime p ==> 0 < p^a"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    30
by (force simp add: prime_iff)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    31
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    32
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
    33
by (rule ccontr, simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    34
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    35
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    36
lemma prime_dvd_cases:
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    37
  fixes p::nat
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    38
  shows "[| p*k dvd m*n;  prime p |]  
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
    39
   ==> (\<exists>x. k dvd x*n & m = p*x) | (\<exists>y. k dvd m*y & n = p*y)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    40
apply (simp add: prime_iff)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    41
apply (frule dvd_mult_left)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    42
apply (subgoal_tac "p dvd m | p dvd n")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    43
 prefer 2 apply blast
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    44
apply (erule disjE)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    45
apply (rule disjI1)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    46
apply (rule_tac [2] disjI2)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27105
diff changeset
    47
apply (auto elim!: dvdE)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    48
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    49
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    50
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    51
lemma prime_power_dvd_cases [rule_format (no_asm)]: 
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    52
fixes p::nat
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    53
  shows "prime p
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
    54
  ==> \<forall>m n. p^c dvd m*n -->  
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
    55
        (\<forall>a b. a+b = Suc c --> p^a dvd m | p^b dvd n)"
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
    56
apply (induct c)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    57
apply (metis dvd_1_left nat_power_eq_Suc_0_iff one_is_add)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    58
(*inductive step*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    59
apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    60
apply clarify
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    61
apply (erule prime_dvd_cases [THEN disjE], assumption, auto)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    62
(*case 1: p dvd m*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    63
 apply (case_tac "a")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    64
  apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    65
 apply clarify
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    66
 apply (drule spec, drule spec, erule (1) notE impE)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    67
 apply (drule_tac x = nat in spec)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    68
 apply (drule_tac x = b in spec)
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
(*case 2: p dvd n*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    71
apply (case_tac "b")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    72
 apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    73
apply clarify
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    74
apply (drule spec, drule spec, erule (1) notE impE)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    75
apply (drule_tac x = a in spec)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    76
apply (drule_tac x = nat in spec, simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    77
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    78
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    79
(*needed in this form in Sylow.ML*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    80
lemma div_combine:
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    81
  fixes p::nat
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
    82
  shows "[| prime p; ~ (p ^ (Suc r) dvd n);  p^(a+r) dvd n*k |] ==> p ^ a dvd k"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55157
diff changeset
    83
by (metis add_Suc add.commute prime_power_dvd_cases)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    84
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    85
(*Lemma for power_dvd_bound*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    86
lemma Suc_le_power: "Suc 0 < p ==> Suc n <= p^n"
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
    87
apply (induct n)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    88
apply (simp (no_asm_simp))
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    89
apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    90
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
    91
apply (subgoal_tac "2 * p^n <= p * p^n")
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
    92
apply arith
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    93
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
    94
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    95
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    96
(*An upper bound for the n such that p^n dvd a: needed for GREATEST to exist*)
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
    97
lemma power_dvd_bound: "[|p^n dvd a;  Suc 0 < p;  a > 0|] ==> n < a"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    98
apply (drule dvd_imp_le)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
    99
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
   100
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   101
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   102
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
   103
text{*Exponent Theorems*}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   104
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   105
lemma exponent_ge [rule_format]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   106
  "[|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
   107
apply (simp add: exponent_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   108
apply (erule Greatest_le)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   109
apply (blast dest: prime_gt_Suc_0_nat power_dvd_bound)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   110
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   111
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   112
lemma power_exponent_dvd: "s>0 ==> (p ^ exponent p s) dvd s"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   113
apply (simp add: exponent_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   114
apply clarify
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   115
apply (rule_tac k = 0 in GreatestI)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   116
prefer 2 apply (blast dest: prime_gt_Suc_0_nat power_dvd_bound, simp)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   117
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   118
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   119
lemma power_Suc_exponent_Not_dvd:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   120
  "[|(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
   121
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
   122
 prefer 2 apply simp 
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   123
apply (rule ccontr)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   124
apply (drule exponent_ge, auto)
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
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   127
lemma exponent_power_eq [simp]: "prime p ==> exponent p (p^a) = a"
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   128
apply (simp add: exponent_def)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   129
apply (rule Greatest_equality, simp)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   130
apply (simp (no_asm_simp) add: prime_gt_Suc_0_nat power_dvd_imp_le)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   131
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   132
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   133
lemma exponent_equalityI:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   134
  "!r::nat. (p^r dvd a) = (p^r dvd b) ==> exponent p a = exponent p b"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   135
by (simp (no_asm_simp) add: exponent_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   136
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   137
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
   138
by (simp (no_asm_simp) add: exponent_def)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   139
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   140
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   141
(* exponent_mult_add, easy inclusion.  Could weaken p \<in> prime to Suc 0 < p *)
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   142
lemma exponent_mult_add1: "[| a > 0; b > 0 |]
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   143
  ==> (exponent p a) + (exponent p b) <= exponent p (a * b)"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   144
apply (case_tac "prime p")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   145
apply (rule exponent_ge)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   146
apply (auto simp add: power_add)
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   147
by (metis mult_dvd_mono power_exponent_dvd)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   148
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   149
(* exponent_mult_add, opposite inclusion *)
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   150
lemma exponent_mult_add2: "[| a > 0; b > 0 |]  
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   151
  ==> exponent p (a * b) <= (exponent p a) + (exponent p b)"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   152
apply (case_tac "prime p")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   153
apply (rule leI, clarify)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   154
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
   155
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
   156
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
   157
  prefer 3 apply assumption
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   158
 prefer 2 apply simp 
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   159
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
   160
 apply (assumption, force, simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   161
apply (blast dest: power_Suc_exponent_Not_dvd)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   162
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   163
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   164
lemma exponent_mult_add: "[| a > 0; b > 0 |]
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   165
   ==> exponent p (a * b) = (exponent p a) + (exponent p b)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   166
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
   167
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   168
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   169
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
   170
apply (case_tac "exponent p n", simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   171
apply (case_tac "n", simp)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   172
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
   173
apply (auto dest: dvd_mult_left)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   174
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   175
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   176
lemma exponent_1_eq_0 [simp]:
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   177
  fixes p::nat
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   178
  shows "exponent p (Suc 0) = 0"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   179
apply (case_tac "prime p")
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   180
apply (metis exponent_power_eq nat_power_eq_Suc_0_iff)
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   181
apply (simp add: prime_iff not_divides_exponent_0)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   182
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   183
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   184
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27651
diff changeset
   185
text{*Main Combinatorial Argument*}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   186
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   187
lemma gcd_mult': fixes a::nat shows "gcd b (a * b) = b"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55157
diff changeset
   188
by (simp add: mult.commute[of a b]) 
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   189
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   190
lemma le_extend_mult: "[| c > 0; a <= b |] ==> a <= b * (c::nat)"
14889
d7711d6b9014 moved some cardinality results into main HOL
paulson
parents: 14706
diff changeset
   191
apply (rule_tac P = "%x. x <= b * c" in subst)
d7711d6b9014 moved some cardinality results into main HOL
paulson
parents: 14706
diff changeset
   192
apply (rule mult_1_right)
d7711d6b9014 moved some cardinality results into main HOL
paulson
parents: 14706
diff changeset
   193
apply (rule mult_le_mono, auto)
d7711d6b9014 moved some cardinality results into main HOL
paulson
parents: 14706
diff changeset
   194
done
d7711d6b9014 moved some cardinality results into main HOL
paulson
parents: 14706
diff changeset
   195
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   196
lemma p_fac_forw_lemma:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   197
  "[| (m::nat) > 0; k > 0; k < p^a; (p^r) dvd (p^a)* m - k |] ==> r <= a"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   198
apply (rule notnotD)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   199
apply (rule notI)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   200
apply (drule contrapos_nn [OF _ leI, THEN notnotD], assumption)
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 23976
diff changeset
   201
apply (drule less_imp_le [of a])
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   202
apply (drule le_imp_power_dvd)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27105
diff changeset
   203
apply (drule_tac b = "p ^ r" in dvd_trans, assumption)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 55157
diff changeset
   204
apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2_nat gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less mult.commute not_add_less2 xt1(10))
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   205
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   206
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   207
lemma p_fac_forw: "[| (m::nat) > 0; k>0; k < p^a; (p^r) dvd (p^a)* m - k |]  
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   208
  ==> (p^r) dvd (p^a) - k"
30011
cc264a9a033d consider changes variable names in theorem le_imp_power_dvd
haftmann
parents: 27717
diff changeset
   209
apply (frule p_fac_forw_lemma [THEN le_imp_power_dvd, of _ k p], auto)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   210
apply (subgoal_tac "p^r dvd p^a*m")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   211
 prefer 2 apply (blast intro: dvd_mult2)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   212
apply (drule dvd_diffD1)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   213
  apply assumption
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31717
diff changeset
   214
 prefer 2 apply (blast intro: dvd_diff_nat)
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   215
apply (drule gr0_implies_Suc, auto)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   216
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   217
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   218
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   219
lemma r_le_a_forw:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   220
  "[| (k::nat) > 0; k < p^a; p>0; (p^r) dvd (p^a) - k |] ==> r <= a"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   221
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
   222
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   223
lemma p_fac_backw: "[| m>0; k>0; (p::nat)\<noteq>0;  k < p^a;  (p^r) dvd p^a - k |]  
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   224
  ==> (p^r) dvd (p^a)*m - k"
30011
cc264a9a033d consider changes variable names in theorem le_imp_power_dvd
haftmann
parents: 27717
diff changeset
   225
apply (frule_tac k1 = k and p1 = p in r_le_a_forw [THEN le_imp_power_dvd], auto)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   226
apply (subgoal_tac "p^r dvd p^a*m")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   227
 prefer 2 apply (blast intro: dvd_mult2)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   228
apply (drule dvd_diffD1)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   229
  apply assumption
31952
40501bb2d57c renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
nipkow
parents: 31717
diff changeset
   230
 prefer 2 apply (blast intro: dvd_diff_nat)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   231
apply (drule less_imp_Suc_add, auto)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   232
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   233
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   234
lemma exponent_p_a_m_k_equation: "[| m>0; k>0; (p::nat)\<noteq>0;  k < p^a |]  
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   235
  ==> exponent p (p^a * m - k) = exponent p (p^a - k)"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   236
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
   237
done
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
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
   240
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
   241
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   242
(*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
   243
lemma p_not_div_choose_lemma [rule_format]:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   244
  "[| \<forall>i. Suc i < K --> exponent p (Suc i) = exponent p (Suc(j+i))|]  
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   245
   ==> k<K --> exponent p ((j+k) choose k) = 0"
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   246
apply (cases "prime p")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   247
 prefer 2 apply simp 
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   248
apply (induct k)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   249
apply (simp (no_asm))
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   250
(*induction step*)
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   251
apply (subgoal_tac "(Suc (j+k) choose Suc k) > 0")
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
   252
 prefer 2 apply (simp, clarify)
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   253
apply (subgoal_tac "exponent p ((Suc (j+k) choose Suc k) * Suc k) = 
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   254
                    exponent p (Suc k)")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   255
 txt{*First, use the assumed equation.  We simplify the LHS to
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 25162
diff changeset
   256
  @{term "exponent p (Suc (j + k) choose Suc k) + exponent p (Suc k)"}
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   257
  the common terms cancel, proving the conclusion.*}
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   258
 apply (simp del: bad_Sucs add: exponent_mult_add)
58917
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   259
apply (simp del: bad_Sucs add: mult_ac Suc_times_binomial exponent_mult_add)
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   260
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   261
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   262
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   263
(*The lemma above, with two changes of variables*)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   264
lemma p_not_div_choose:
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   265
  "[| k<K;  k<=n;
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   266
      \<forall>j. 0<j & j<K --> exponent p (n - k + (K - j)) = exponent p (K - j)|]
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 24742
diff changeset
   267
   ==> exponent p (n choose k) = 0"
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   268
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
   269
  prefer 3 apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   270
 prefer 2 apply assumption
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   271
apply (drule_tac x = "K - Suc i" in spec)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   272
apply (simp add: Suc_diff_le)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   273
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   274
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   275
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   276
lemma const_p_fac_right:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   277
  "m>0 ==> exponent p ((p^a * m - Suc 0) choose (p^a - Suc 0)) = 0"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   278
apply (case_tac "prime p")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   279
 prefer 2 apply simp 
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   280
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
   281
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
   282
   apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   283
  apply simp
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   284
 apply (case_tac "m")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   285
  apply (case_tac [2] "p^a")
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   286
   apply auto
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   287
(*now the hard case, simplified to
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   288
    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
   289
apply (subgoal_tac "0<p")
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   290
 prefer 2 apply (force dest!: prime_gt_Suc_0_nat)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   291
apply (subst exponent_p_a_m_k_equation, auto)
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
lemma const_p_fac:
25162
ad4d5365d9d8 went back to >0
nipkow
parents: 25134
diff changeset
   295
  "m>0 ==> exponent p (((p^a) * m) choose p^a) = exponent p m"
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 16417
diff changeset
   296
apply (case_tac "prime p")
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   297
 prefer 2 apply simp 
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   298
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
   299
 prefer 2 apply (force simp add: prime_iff)
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   300
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
   301
  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
   302
  first
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   303
  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
   304
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
   305
                    a + exponent p m")
55157
06897ea77f78 converting to new Number_Theory
paulson <lp15@cam.ac.uk>
parents: 41413
diff changeset
   306
 apply (simp add: exponent_mult_add)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   307
txt{*one subgoal left!*}
58917
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   308
apply (auto simp: mult_ac)
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   309
apply (subst times_binomial_minus1_eq, simp)
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   310
apply (simp add: diff_le_mono exponent_mult_add)
a3be9a47e2d7 Tidying up. Removing unnecessary conditions from some theorems.
paulson <lp15@cam.ac.uk>
parents: 57865
diff changeset
   311
apply (metis const_p_fac_right mult.commute)
13870
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   312
done
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   313
cf947d1ec5ff moved Exponent, Coset, Sylow from GroupTheory to Algebra, converting them
paulson
parents:
diff changeset
   314
end