]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/arithmetics/factorial.ma
fc2320a948cbd324ddfe900db2af3f33837866f7
[helm.git] / matita / matita / lib / arithmetics / factorial.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic        
3     ||A||  Library of Mathematics, developed at the Computer Science     
4     ||T||  Department of the University of Bologna, Italy.                     
5     ||I||                                                                 
6     ||T||  
7     ||A||  This file is distributed under the terms of the 
8     \   /  GNU General Public License Version 2        
9      \ /      
10       V_______________________________________________________________ *)
11
12 include "arithmetics/exp.ma".
13
14 let rec fact n ≝
15   match n with 
16   [ O ⇒ 1
17   | S m ⇒ fact m * S m].
18   
19 interpretation "factorial" 'fact n = (fact n).
20
21 lemma factS: ∀n. (S n)! = (S n)*n!.
22 #n >commutative_times // qed.
23
24 theorem le_1_fact : ∀n. 1 ≤ n!.
25 #n (elim n) normalize /2 by lt_minus_to_plus/ 
26 qed.
27
28 theorem le_2_fact : ∀n. 1 < n → 2 ≤ n!.
29 #n (cases n)
30   [#abs @False_ind /2/
31   |#m normalize #le2 @(le_times 1 ? 2) //
32   ]
33 qed.
34
35 theorem le_n_fact_n: ∀n. n ≤ n!.
36 #n (elim n) normalize //
37 #n #Hind @(transitive_le ? (1*(S n))) // @le_times //
38 qed.
39
40 theorem lt_n_fact_n: ∀n. 2 < n → n < n!.
41 #n (cases n) 
42   [#H @False_ind /2/ 
43   |#m #lt2 normalize @(lt_to_le_to_lt ? (2*(S m))) //
44    @le_times // @le_2_fact /2 by lt_plus_to_lt_l/ 
45 qed.
46
47 (* approximations *)
48
49 theorem fact_to_exp1: ∀n.O<n →
50  (2*n)! ≤ (2^(pred (2*n))) * n! * n!.
51 #n #posn (cut (∀i.2*(S i) = S(S(2*i)))) [//] #H (elim posn) //
52 #n #posn #Hind @(transitive_le ? ((2*n)!*(2*(S n))*(2*(S n))))
53   [>H normalize @le_times //
54   |cut (pred (2*(S n)) = S(S(pred(2*n))))
55     [>S_pred // @(le_times 1 ? 1) //] #H1
56    cut (2^(pred (2*(S n))) = 2^(pred(2*n))*2*2) 
57     [>H1 >H1 //] #H2 >H2
58    @(transitive_le ? ((2^(pred (2*n))) * n! * n! *(2*(S n))*(2*(S n))))
59     [@le_times[@le_times //]//
60     (* we generalize to hide the computational content *)
61     |normalize in match ((S n)!); generalize in match (S n);
62      #Sn generalize in match 2; #two //
63     ]
64   ]
65 qed.  
66
67 theorem fact_to_exp: ∀n.
68  (2*n)! ≤ (2^(pred (2*n))) * n! * n!.
69 #n (cases n) [normalize // |#n @fact_to_exp1 //]
70 qed.
71
72 theorem exp_to_fact1: ∀n.O < n →
73   2^(2*n)*n!*n! < (S(2*n))!.
74 #n #posn (elim posn) [normalize //]
75 #n #posn #Hind (cut (∀i.2*(S i) = S(S(2*i)))) [//] #H
76 cut (2^(2*(S n)) = 2^(2*n)*2*2) [>H //] #H1 >H1
77 @(le_to_lt_to_lt ? (2^(2*n)*n!*n!*(2*(S n))*(2*(S n))))
78   [normalize in match ((S n)!); generalize in match (S n); #Sn
79    generalize in match 2; #two //
80   |cut ((S(2*(S n)))! = (S(2*n))!*(S(S(2*n)))*(S(S(S(2*n)))))
81    [>H //] #H2 >H2 @lt_to_le_to_lt_times
82      [@lt_to_le_to_lt_times //|>H // | //]
83   ]
84 qed.
85
86 (* a sligtly better result *)
87 theorem exp_to_fact2: ∀n.O < n →
88   (exp 2 (2*n))*(exp (fact n) 2) \le 2*n*fact (2*n).
89 #n #posn elim posn
90   [@le_n
91   |#m #le1m #Hind 
92    cut (2*(S m) = S(S (2*m))) [normalize //] #H2 >H2 in ⊢ (?%?);
93    >factS <times_exp 
94    whd in match (exp 2 (S(S ?))); >(commutative_times ? 2) >associative_times
95    >associative_times in ⊢ (??%); @monotonic_le_times_r 
96    whd in match (exp 2 (S ?)); >(commutative_times ? 2) >associative_times
97    @(transitive_le ? (2*((2*m*(2*m)!)*(S m)^2)))
98     [@le_times [//] >commutative_times in ⊢ (?(??%)?); <associative_times
99      @le_times [@Hind |@le_n] 
100     |>exp_2 <associative_times <associative_times >commutative_times in ⊢ (??%);
101      @le_times [2:@le_n] >H2 >factS >commutative_times <associative_times 
102      @le_times //
103     ]
104   ]
105 qed.
106
107 theorem le_fact_10: fact (2*5) ≤ (exp 2 ((2*5)-2))*(fact 5)*(fact 5).
108 >factS in ⊢ (?%?); >factS in ⊢ (?%?); <associative_times in ⊢ (?%?);
109 >factS in ⊢ (?%?); <associative_times in ⊢ (?%?);
110 >factS in ⊢ (?%?); <associative_times in ⊢ (?%?);
111 >factS in ⊢ (?%?); <associative_times in ⊢ (?%?);
112 @le_times [2:%] @leb_true_to_le % 
113 qed-.
114
115 theorem ab_times_cd: ∀a,b,c,d.(a*b)*(c*d)=(a*c)*(b*d). 
116 //
117 qed.
118
119 (* an even better result *)
120 theorem lt_4_to_fact: ∀n.4<n →
121   fact (2*n) ≤ (exp 2 ((2*n)-2))*(fact n)*(fact n).
122 #n #ltn elim ltn
123   [@le_fact_10
124   |#m #lem #Hind 
125    cut (2*(S m) = S(S (2*m))) [normalize //] #H2 >H2 
126    whd in match (minus (S(S ?)) 2); <minus_n_O
127    >factS >factS <associative_times >factS
128    @(transitive_le ? ((2*(S m))*(2*(S m))*(fact (2*m))))
129     [@le_times [2:@le_n] >H2 @le_times //
130     |@(transitive_le ? (2*S m*(2*S m)*(2\sup(2*m-2)*m!*m!)))
131       [@monotonic_le_times_r //
132       |>associative_times >ab_times_cd in ⊢ (?(??%)?);
133        <associative_times @le_times [2:@le_n] 
134        <associative_times in ⊢ (?(??%)?);
135        >ab_times_cd @le_times [2:@le_n] >commutative_times
136        >(commutative_times 2) @(le_exp (S(S ((2*m)-2)))) [//]
137        >eq_minus_S_pred >S_pred 
138         [>eq_minus_S_pred >S_pred [<minus_n_O @le_n |elim lem //]
139         |elim lem [//] #m0 #le5m0 #Hind 
140          normalize <plus_n_Sm //
141         ]
142       ]
143     ]
144   ]
145 qed.