src/HOL/Fact.thy
author hoelzl
Tue, 27 Nov 2012 10:56:31 +0100
changeset 50240 019d642d422d
parent 50224 aacd6da09825
child 57113 7e95523302e6
permissions -rw-r--r--
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
     1
(*  Title       : Fact.thy
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
     5
    The integer version of factorial and other additions by Jeremy Avigad.
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     6
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     7
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
     8
header{*Factorial Function*}
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15094
diff changeset
    10
theory Fact
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32558
diff changeset
    11
imports Main
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15094
diff changeset
    12
begin
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
    13
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    14
class fact =
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
    15
  fixes fact :: "'a \<Rightarrow> 'a"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    16
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    17
instantiation nat :: fact
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    18
begin 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    19
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    20
fun
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    21
  fact_nat :: "nat \<Rightarrow> nat"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    22
where
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    23
  fact_0_nat: "fact_nat 0 = Suc 0"
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32039
diff changeset
    24
| fact_Suc: "fact_nat (Suc x) = Suc x * fact x"
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    25
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
    26
instance ..
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    27
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    28
end
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    29
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    30
(* definitions for the integers *)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    31
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    32
instantiation int :: fact
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    33
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    34
begin 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    35
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    36
definition
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    37
  fact_int :: "int \<Rightarrow> int"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    38
where  
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    39
  "fact_int x = (if x >= 0 then int (fact (nat x)) else 0)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    40
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    41
instance proof qed
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    42
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    43
end
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    44
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    45
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    46
subsection {* Set up Transfer *}
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    47
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    48
lemma transfer_nat_int_factorial:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    49
  "(x::int) >= 0 \<Longrightarrow> fact (nat x) = nat (fact x)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    50
  unfolding fact_int_def
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    51
  by auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    52
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    53
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    54
lemma transfer_nat_int_factorial_closure:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    55
  "x >= (0::int) \<Longrightarrow> fact x >= 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    56
  by (auto simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    57
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35028
diff changeset
    58
declare transfer_morphism_nat_int[transfer add return: 
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    59
    transfer_nat_int_factorial transfer_nat_int_factorial_closure]
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    60
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    61
lemma transfer_int_nat_factorial:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    62
  "fact (int x) = int (fact x)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    63
  unfolding fact_int_def by auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    64
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    65
lemma transfer_int_nat_factorial_closure:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    66
  "is_nat x \<Longrightarrow> fact x >= 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    67
  by (auto simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    68
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35028
diff changeset
    69
declare transfer_morphism_int_nat[transfer add return: 
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    70
    transfer_int_nat_factorial transfer_int_nat_factorial_closure]
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
    71
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
    72
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    73
subsection {* Factorial *}
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    74
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    75
lemma fact_0_int [simp]: "fact (0::int) = 1"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    76
  by (simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    77
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    78
lemma fact_1_nat [simp]: "fact (1::nat) = 1"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    79
  by simp
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    80
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    81
lemma fact_Suc_0_nat [simp]: "fact (Suc 0) = Suc 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    82
  by simp
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    83
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    84
lemma fact_1_int [simp]: "fact (1::int) = 1"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    85
  by (simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    86
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    87
lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    88
  by simp
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    89
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    90
lemma fact_plus_one_int: 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    91
  assumes "n >= 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    92
  shows "fact ((n::int) + 1) = (n + 1) * fact n"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
    93
  using assms unfolding fact_int_def 
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    94
  by (simp add: nat_add_distrib algebra_simps int_mult)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    95
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    96
lemma fact_reduce_nat: "(n::nat) > 0 \<Longrightarrow> fact n = n * fact (n - 1)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    97
  apply (subgoal_tac "n = Suc (n - 1)")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
    98
  apply (erule ssubst)
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32039
diff changeset
    99
  apply (subst fact_Suc)
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   100
  apply simp_all
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   101
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   102
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   103
lemma fact_reduce_int: "(n::int) > 0 \<Longrightarrow> fact n = n * fact (n - 1)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   104
  apply (subgoal_tac "n = (n - 1) + 1")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   105
  apply (erule ssubst)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   106
  apply (subst fact_plus_one_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   107
  apply simp_all
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   108
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   109
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   110
lemma fact_nonzero_nat [simp]: "fact (n::nat) \<noteq> 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   111
  apply (induct n)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   112
  apply (auto simp add: fact_plus_one_nat)
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   113
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   114
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   115
lemma fact_nonzero_int [simp]: "n >= 0 \<Longrightarrow> fact (n::int) ~= 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   116
  by (simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   117
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   118
lemma fact_gt_zero_nat [simp]: "fact (n :: nat) > 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   119
  by (insert fact_nonzero_nat [of n], arith)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   120
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   121
lemma fact_gt_zero_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) > 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   122
  by (auto simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   123
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   124
lemma fact_ge_one_nat [simp]: "fact (n :: nat) >= 1"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   125
  by (insert fact_nonzero_nat [of n], arith)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   126
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   127
lemma fact_ge_Suc_0_nat [simp]: "fact (n :: nat) >= Suc 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   128
  by (insert fact_nonzero_nat [of n], arith)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   129
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   130
lemma fact_ge_one_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) >= 1"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   131
  apply (auto simp add: fact_int_def)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   132
  apply (subgoal_tac "1 = int 1")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   133
  apply (erule ssubst)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   134
  apply (subst zle_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   135
  apply auto
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   136
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   137
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   138
lemma dvd_fact_nat [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::nat)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   139
  apply (induct n)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   140
  apply force
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32039
diff changeset
   141
  apply (auto simp only: fact_Suc)
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   142
  apply (subgoal_tac "m = Suc n")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   143
  apply (erule ssubst)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   144
  apply (rule dvd_triv_left)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   145
  apply auto
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   146
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   147
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   148
lemma dvd_fact_int [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::int)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   149
  apply (case_tac "1 <= n")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   150
  apply (induct n rule: int_ge_induct)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   151
  apply (auto simp add: fact_plus_one_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   152
  apply (subgoal_tac "m = i + 1")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   153
  apply auto
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 40033
diff changeset
   154
  done
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   155
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   156
lemma interval_plus_one_nat: "(i::nat) <= j + 1 \<Longrightarrow> 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   157
  {i..j+1} = {i..j} Un {j+1}"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   158
  by auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   159
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   160
lemma interval_Suc: "i <= Suc j \<Longrightarrow> {i..Suc j} = {i..j} Un {Suc j}"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   161
  by auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   162
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   163
lemma interval_plus_one_int: "(i::int) <= j + 1 \<Longrightarrow> {i..j+1} = {i..j} Un {j+1}"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   164
  by auto
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   165
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   166
lemma fact_altdef_nat: "fact (n::nat) = (PROD i:{1..n}. i)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   167
  apply (induct n)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   168
  apply force
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32039
diff changeset
   169
  apply (subst fact_Suc)
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   170
  apply (subst interval_Suc)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   171
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   172
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   173
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   174
lemma fact_altdef_int: "n >= 0 \<Longrightarrow> fact (n::int) = (PROD i:{1..n}. i)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   175
  apply (induct n rule: int_ge_induct)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   176
  apply force
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   177
  apply (subst fact_plus_one_int, assumption)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   178
  apply (subst interval_plus_one_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   179
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   180
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   181
40033
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   182
lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd fact (m::nat)"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   183
  by (auto simp add: fact_altdef_nat intro!: setprod_dvd_setprod_subset)
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   184
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   185
lemma fact_mod: "m \<le> (n::nat) \<Longrightarrow> fact n mod fact m = 0"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   186
  by (auto simp add: dvd_imp_mod_0 fact_dvd)
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   187
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   188
lemma fact_div_fact:
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   189
  assumes "m \<ge> (n :: nat)"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   190
  shows "(fact m) div (fact n) = \<Prod>{n + 1..m}"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   191
proof -
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   192
  obtain d where "d = m - n" by auto
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   193
  from assms this have "m = n + d" by auto
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   194
  have "fact (n + d) div (fact n) = \<Prod>{n + 1..n + d}"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   195
  proof (induct d)
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   196
    case 0
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   197
    show ?case by simp
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   198
  next
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   199
    case (Suc d')
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   200
    have "fact (n + Suc d') div fact n = Suc (n + d') * fact (n + d') div fact n"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   201
      by simp
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   202
    also from Suc.hyps have "... = Suc (n + d') * \<Prod>{n + 1..n + d'}" 
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   203
      unfolding div_mult1_eq[of _ "fact (n + d')"] by (simp add: fact_mod)
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   204
    also have "... = \<Prod>{n + 1..n + Suc d'}"
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   205
      by (simp add: atLeastAtMostSuc_conv setprod_insert)
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   206
    finally show ?case .
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   207
  qed
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   208
  from this `m = n + d` show ?thesis by simp
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   209
qed
84200d970bf0 added some facts about factorial and dvd, div and mod
bulwahn
parents: 35644
diff changeset
   210
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   211
lemma fact_mono_nat: "(m::nat) \<le> n \<Longrightarrow> fact m \<le> fact n"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   212
apply (drule le_imp_less_or_eq)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   213
apply (auto dest!: less_imp_Suc_add)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   214
apply (induct_tac k, auto)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   215
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   216
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   217
lemma fact_neg_int [simp]: "m < (0::int) \<Longrightarrow> fact m = 0"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   218
  unfolding fact_int_def by auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   219
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   220
lemma fact_ge_zero_int [simp]: "fact m >= (0::int)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   221
  apply (case_tac "m >= 0")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   222
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   223
  apply (frule fact_gt_zero_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   224
  apply arith
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   225
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   226
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   227
lemma fact_mono_int_aux [rule_format]: "k >= (0::int) \<Longrightarrow> 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   228
    fact (m + k) >= fact m"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   229
  apply (case_tac "m < 0")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   230
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   231
  apply (induct k rule: int_ge_induct)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   232
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   233
  apply (subst add_assoc [symmetric])
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   234
  apply (subst fact_plus_one_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   235
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   236
  apply (erule order_trans)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   237
  apply (subst mult_le_cancel_right1)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   238
  apply (subgoal_tac "fact (m + i) >= 0")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   239
  apply arith
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   240
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   241
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   242
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   243
lemma fact_mono_int: "(m::int) <= n \<Longrightarrow> fact m <= fact n"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   244
  apply (insert fact_mono_int_aux [of "n - m" "m"])
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   245
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   246
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   247
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   248
text{*Note that @{term "fact 0 = fact 1"}*}
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   249
lemma fact_less_mono_nat: "[| (0::nat) < m; m < n |] ==> fact m < fact n"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   250
apply (drule_tac m = m in less_imp_Suc_add, auto)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   251
apply (induct_tac k, auto)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   252
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   253
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   254
lemma fact_less_mono_int_aux: "k >= 0 \<Longrightarrow> (0::int) < m \<Longrightarrow>
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   255
    fact m < fact ((m + 1) + k)"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   256
  apply (induct k rule: int_ge_induct)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   257
  apply (simp add: fact_plus_one_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   258
  apply (subst (2) fact_reduce_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   259
  apply (auto simp add: add_ac)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   260
  apply (erule order_less_le_trans)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   261
  apply (subst mult_le_cancel_right1)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   262
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   263
  apply (subgoal_tac "fact (i + (1 + m)) >= 0")
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   264
  apply force
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   265
  apply (rule fact_ge_zero_int)
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   266
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   267
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   268
lemma fact_less_mono_int: "(0::int) < m \<Longrightarrow> m < n \<Longrightarrow> fact m < fact n"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   269
  apply (insert fact_less_mono_int_aux [of "n - (m + 1)" "m"])
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   270
  apply auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   271
done
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   272
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   273
lemma fact_num_eq_if_nat: "fact (m::nat) = 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   274
  (if m=0 then 1 else m * fact (m - 1))"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   275
by (cases m) auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   276
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   277
lemma fact_add_num_eq_if_nat:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   278
  "fact ((m::nat) + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   279
by (cases "m + n") auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   280
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   281
lemma fact_add_num_eq_if2_nat:
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   282
  "fact ((m::nat) + n) = 
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   283
    (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))"
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   284
by (cases m) auto
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   285
45930
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   286
lemma fact_le_power: "fact n \<le> n^n"
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   287
proof (induct n)
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   288
  case (Suc n)
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   289
  then have "fact n \<le> Suc n ^ n" by (rule le_trans) (simp add: power_mono)
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   290
  then show ?case by (simp add: add_le_mono)
2a882ef2cd73 add lemmas
noschinl
parents: 41550
diff changeset
   291
qed simp
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 30242
diff changeset
   292
32039
400a519bc888 Use term antiquotation to refer to constant names in subsection title.
berghofe
parents: 32036
diff changeset
   293
subsection {* @{term fact} and @{term of_nat} *}
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   294
29693
708dcf7dec9f moved upwards in thy graph, real related theorems moved to Transcendental.thy
chaieb
parents: 28952
diff changeset
   295
lemma of_nat_fact_not_zero [simp]: "of_nat (fact n) \<noteq> (0::'a::semiring_char_0)"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   296
by auto
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   297
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33319
diff changeset
   298
lemma of_nat_fact_gt_zero [simp]: "(0::'a::{linordered_semidom}) < of_nat(fact n)" by auto
29693
708dcf7dec9f moved upwards in thy graph, real related theorems moved to Transcendental.thy
chaieb
parents: 28952
diff changeset
   299
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33319
diff changeset
   300
lemma of_nat_fact_ge_zero [simp]: "(0::'a::linordered_semidom) \<le> of_nat(fact n)"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   301
by simp
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   302
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33319
diff changeset
   303
lemma inv_of_nat_fact_gt_zero [simp]: "(0::'a::linordered_field) < inverse (of_nat (fact n))"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   304
by (auto simp add: positive_imp_inverse_positive)
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   305
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33319
diff changeset
   306
lemma inv_of_nat_fact_ge_zero [simp]: "(0::'a::linordered_field) \<le> inverse (of_nat (fact n))"
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 25112
diff changeset
   307
by (auto intro: order_less_imp_le)
15094
a7d1a3fdc30d conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents: 12196
diff changeset
   308
50224
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   309
lemma fact_eq_rev_setprod_nat: "fact (k::nat) = (\<Prod>i<k. k - i)"
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   310
  unfolding fact_altdef_nat
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   311
proof (rule strong_setprod_reindex_cong)
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   312
  { fix i assume "Suc 0 \<le> i" "i \<le> k" then have "\<exists>j\<in>{..<k}. i = k - j"
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   313
      by (intro bexI[of _ "k - i"]) simp_all }
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   314
  then show "{1..k} = (\<lambda>i. k - i) ` {..<k}"
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   315
    by (auto simp: image_iff)
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   316
qed (auto intro: inj_onI)
aacd6da09825 add binomial_ge_n_over_k_pow_k
hoelzl
parents: 46240
diff changeset
   317
50240
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   318
lemma fact_div_fact_le_pow:
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   319
  assumes "r \<le> n" shows "fact n div fact (n - r) \<le> n ^ r"
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   320
proof -
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   321
  have "\<And>r. r \<le> n \<Longrightarrow> \<Prod>{n - r..n} = (n - r) * \<Prod>{Suc (n - r)..n}"
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   322
    by (subst setprod_insert[symmetric]) (auto simp: atLeastAtMost_insertL)
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   323
  with assms show ?thesis
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   324
    by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono)
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   325
qed
019d642d422d add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents: 50224
diff changeset
   326
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15094
diff changeset
   327
end