]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/arithmetics/exp.ma
654185c89cb014f498e0bff30a9a0b63b8a4333a
[helm.git] / matita / matita / lib / arithmetics / exp.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/div_and_mod.ma".
13
14 let rec exp n m on m ≝ 
15  match m with 
16  [ O ⇒ 1
17  | S p ⇒ (exp n p) * n].
18
19 interpretation "natural exponent" 'exp a b = (exp a b).
20
21 theorem exp_plus_times : ∀n,p,q:nat. 
22   n^(p + q) = n^p * n^q.
23 #n #p #q elim p normalize //
24 qed.
25
26 theorem exp_n_O : ∀n:nat. 1 = n^O. 
27 //
28 qed.
29
30 theorem exp_n_1 : ∀n:nat. n = n^1. 
31 #n normalize //
32 qed.
33
34 theorem exp_1_n : ∀n:nat. 1 = 1^n.
35 #n (elim n) normalize //
36 qed.
37
38 theorem exp_2: ∀n. n^2 = n*n. 
39 #n normalize //
40 qed.
41
42 theorem exp_exp_times : ∀n,p,q:nat. 
43   (n^p)^q = n^(p * q).
44 #n #p #q (elim q) normalize 
45 (* [applyS exp_n_O funziona ma non lo trova *)
46 // <times_n_O // 
47 qed.
48
49 theorem lt_O_exp: ∀n,m:nat. O < n → O < n^m. 
50 #n #m (elim m) normalize // #a #Hind #posn 
51 @(le_times 1 ? 1) /2/ qed.
52
53 (* [applyS monotonic_le_minus_r /2/ 
54
55 > (minus_Sn_n ?) 
56 [cut (∃x.∃y.∃z.(x - y ≤ x - z) = (1 ≤ n ^a))
57   [@ex_intro
58     [|@ex_intro
59       [|@ex_intro
60         [|     
61 @(rewrite_r \Nopf  (S  n \sup a - n \sup a )
62  (\lambda x:\Nopf 
63   .(S  n \sup a - n \sup a \le S  n \sup a -(S ?-?))=(x\le  n \sup a ))
64  (rewrite_r \Nopf  ?
65   (\lambda x:\Nopf 
66    .(S  n \sup a - n \sup a \le x)=(S  n \sup a - n \sup a \le  n \sup a ))
67   ( refl … Type[0]  (S  n \sup a - n \sup a \le  n \sup a ) ) (S ?-(S ?-?))
68   (rewrite_r \Nopf  (?-O) (\lambda x:\Nopf .S ?-(S ?-?)=x)
69    (rewrite_l \Nopf  1 (\lambda x:\Nopf .S ?-x=n ^a -O) (minus_S_S ? O) (S ?-?)
70     (minus_Sn_n ?)) ?
71    (minus_n_O ?))) 1
72  (minus_Sn_n  n \sup a ))  
73          
74 @(rewrite_r \Nopf  (S ?-?)
75  (\lambda x:\Nopf 
76   .(S  n \sup a - n \sup a \le S  n \sup a -(S ?-?))=(x\le  n \sup a ))
77  (rewrite_r \Nopf  ?
78   (\lambda x:\Nopf 
79    .(S  n \sup a - n \sup a \le x)=(S  n \sup a - n \sup a \le  n \sup a ))
80   ( refl ??) (S ?-(S ?-?))
81   ?) 1
82  (minus_Sn_n ?))
83   [||
84 @(rewrite_r \Nopf  (?-O) (\lambda x:\Nopf .S ?-(S ?-?)=x)
85   ???)
86 [|<(minus_Sn_n ?)
87 @(rewrite_r \Nopf  (?-O) (\lambda x:\Nopf .S ?-(S ?-?)=x)
88  (rewrite_l \Nopf  1 (\lambda x:\Nopf .S ?-x=?-O) 
89     ? (S ?-?)
90     (minus_Sn_n ?)) ??)
91  
92 @(rewrite_r \Nopf  (?-O) (\lambda x:\Nopf .S ?-(S ?-?)=x)
93    (rewrite_l \Nopf  1 (\lambda x:\Nopf .S ?-x=?-O) (minus_S_S ? O) (S ?-?)
94     (minus_Sn_n ?)) ?
95    (minus_n_O ?)))
96
97 applyS monotonic_le_minus_r /2/
98 qed. *)
99
100 theorem lt_m_exp_nm: ∀n,m:nat. 1 < n → m < n^m.
101 #n #m #lt1n (elim m) normalize // 
102 #n #Hind @(transitive_le ? ((S n)*2)) // @le_times //
103 qed.
104
105 theorem exp_to_eq_O: ∀n,m:nat. 1 < n → 
106   n^m = 1 → m = O.
107 #n #m #ltin #eq1 @le_to_le_to_eq //
108 @le_S_S_to_le <eq1 @lt_m_exp_nm //
109 qed.
110
111 theorem injective_exp_r: ∀b:nat. 1 < b → 
112   injective nat nat (λi:nat. b^i).
113 #b #lt1b @nat_elim2 normalize 
114   [#n #H @sym_eq @(exp_to_eq_O b n lt1b) //
115   |#n #H @False_ind @(absurd (1 < 1) ? (not_le_Sn_n 1))
116    <H in ⊢ (??%); @(lt_to_le_to_lt ? (1*b)) //
117    @le_times // @lt_O_exp /2/
118   |#n #m #Hind #H @eq_f @Hind @(injective_times_l … H) /2/
119   ]
120 qed.
121
122 theorem le_exp: ∀n,m,p:nat. O < p →
123   n ≤m → p^n ≤ p^m.
124 @nat_elim2 #n #m
125   [#ltm #len @lt_O_exp //
126   |#_ #len @False_ind /2/
127   |#Hind #p #posp #lenm normalize @le_times // @Hind /2/
128   ]
129 qed.
130
131 theorem le_exp1: ∀n,m,a:nat. O < a →
132   n ≤m → n^a ≤ m^a.
133 #n #m #a #posa #lenm (elim posa) //
134 #a #posa #Hind @le_times //
135 qed.
136
137 theorem lt_exp: ∀n,m,p:nat. 1 < p → 
138   n < m → p^n < p^m.
139 #n #m #p #lt1p #ltnm 
140 cut (p \sup n ≤ p \sup m) [@le_exp /2/] #H 
141 cases(le_to_or_lt_eq … H) // #eqexp
142 @False_ind @(absurd (n=m)) /2/
143 qed.
144
145 theorem lt_exp1: ∀n,m,p:nat. 0 < p → 
146   n < m → n^p < m^p.
147 #n #m #p #posp #ltnm (elim posp) //
148 #p #posp #Hind @lt_times //
149 qed.
150   
151 theorem le_exp_to_le: 
152 ∀b,n,m. 1 < b → b^n ≤ b^m → n ≤ m.
153 #b #n #m #lt1b #leexp cases(decidable_le n m) //
154 #notle @False_ind @(absurd … leexp) @lt_to_not_le
155 @lt_exp /2/
156 qed.
157
158 theorem le_exp_to_le1 : ∀n,m,p.O < p → 
159   n^p ≤ m^p → n ≤ m.
160 #n #m #p #posp #leexp @not_lt_to_le 
161 @(not_to_not … (lt_exp1 ??? posp)) @le_to_not_lt // 
162 qed.
163      
164 theorem lt_exp_to_lt: 
165 ∀a,n,m. 0 < a → a^n < a^m → n < m.
166 #a #n #m #lt1a #ltexp cases(decidable_le (S n) m) //
167 #H @False_ind @(absurd … ltexp) @le_to_not_lt 
168 @le_exp // @not_lt_to_le @H
169 qed.
170
171 theorem lt_exp_to_lt1: 
172 ∀a,n,m. O < a → n^a < m^a → n < m.
173 #a #n #m #posa #ltexp cases(decidable_le (S n) m) //
174 #H @False_ind @(absurd … ltexp) @le_to_not_lt 
175 @le_exp1 // @not_lt_to_le @H 
176 qed.
177      
178 theorem times_exp: ∀n,m,p. 
179   n^p * m^p = (n*m)^p.
180 #n #m #p (elim p) // #p #Hind normalize //
181 qed.
182
183   
184    
185