| author | haftmann | 
| Tue, 20 Dec 2016 15:39:13 +0100 | |
| changeset 64630 | 96015aecfeba | 
| parent 64246 | 15d1ee6e847b | 
| child 67613 | ce654b0e6d69 | 
| permissions | -rw-r--r-- | 
| 1476 | 1  | 
(* Title: HOL/Hoare/Arith2.thy  | 
2  | 
Author: Norbert Galm  | 
|
| 1335 | 3  | 
Copyright 1995 TUM  | 
4  | 
||
| 
3372
 
6e472c8f0011
Replacement of "divides" by "dvd" from Divides.thy, and updating of proofs
 
paulson 
parents: 
1824 
diff
changeset
 | 
5  | 
More arithmetic. Much of this duplicates ex/Primes.  | 
| 1335 | 6  | 
*)  | 
7  | 
||
| 17278 | 8  | 
theory Arith2  | 
9  | 
imports Main  | 
|
10  | 
begin  | 
|
| 1335 | 11  | 
|
| 
57442
 
2373b4c61111
removed obsolete Isar system commands -- raw ML console is normally used for system tinkering;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
12  | 
definition cd :: "[nat, nat, nat] => bool"  | 
| 38353 | 13  | 
where "cd x m n \<longleftrightarrow> x dvd m & x dvd n"  | 
| 1335 | 14  | 
|
| 38353 | 15  | 
definition gcd :: "[nat, nat] => nat"  | 
16  | 
where "gcd m n = (SOME x. cd x m n & (!y.(cd y m n) --> y<=x))"  | 
|
| 5183 | 17  | 
|
| 38353 | 18  | 
primrec fac :: "nat => nat"  | 
19  | 
where  | 
|
| 5183 | 20  | 
"fac 0 = Suc 0"  | 
| 38353 | 21  | 
| "fac (Suc n) = Suc n * fac n"  | 
| 1335 | 22  | 
|
| 19802 | 23  | 
|
| 62042 | 24  | 
subsubsection \<open>cd\<close>  | 
| 19802 | 25  | 
|
26  | 
lemma cd_nnn: "0<n ==> cd n n n"  | 
|
27  | 
apply (simp add: cd_def)  | 
|
28  | 
done  | 
|
29  | 
||
30  | 
lemma cd_le: "[| cd x m n; 0<m; 0<n |] ==> x<=m & x<=n"  | 
|
31  | 
apply (unfold cd_def)  | 
|
32  | 
apply (blast intro: dvd_imp_le)  | 
|
33  | 
done  | 
|
34  | 
||
35  | 
lemma cd_swap: "cd x m n = cd x n m"  | 
|
36  | 
apply (unfold cd_def)  | 
|
37  | 
apply blast  | 
|
38  | 
done  | 
|
39  | 
||
40  | 
lemma cd_diff_l: "n<=m ==> cd x m n = cd x (m-n) n"  | 
|
41  | 
apply (unfold cd_def)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
42  | 
apply (fastforce dest: dvd_diffD)  | 
| 19802 | 43  | 
done  | 
44  | 
||
45  | 
lemma cd_diff_r: "m<=n ==> cd x m n = cd x m (n-m)"  | 
|
46  | 
apply (unfold cd_def)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
47  | 
apply (fastforce dest: dvd_diffD)  | 
| 19802 | 48  | 
done  | 
49  | 
||
50  | 
||
| 62042 | 51  | 
subsubsection \<open>gcd\<close>  | 
| 19802 | 52  | 
|
53  | 
lemma gcd_nnn: "0<n ==> n = gcd n n"  | 
|
54  | 
apply (unfold gcd_def)  | 
|
55  | 
apply (frule cd_nnn)  | 
|
56  | 
apply (rule some_equality [symmetric])  | 
|
57  | 
apply (blast dest: cd_le)  | 
|
| 33657 | 58  | 
apply (blast intro: le_antisym dest: cd_le)  | 
| 19802 | 59  | 
done  | 
60  | 
||
61  | 
lemma gcd_swap: "gcd m n = gcd n m"  | 
|
62  | 
apply (simp add: gcd_def cd_swap)  | 
|
63  | 
done  | 
|
64  | 
||
65  | 
lemma gcd_diff_l: "n<=m ==> gcd m n = gcd (m-n) n"  | 
|
66  | 
apply (unfold gcd_def)  | 
|
67  | 
apply (subgoal_tac "n<=m ==> !x. cd x m n = cd x (m-n) n")  | 
|
68  | 
apply simp  | 
|
69  | 
apply (rule allI)  | 
|
70  | 
apply (erule cd_diff_l)  | 
|
71  | 
done  | 
|
72  | 
||
73  | 
lemma gcd_diff_r: "m<=n ==> gcd m n = gcd m (n-m)"  | 
|
74  | 
apply (unfold gcd_def)  | 
|
75  | 
apply (subgoal_tac "m<=n ==> !x. cd x m n = cd x m (n-m) ")  | 
|
76  | 
apply simp  | 
|
77  | 
apply (rule allI)  | 
|
78  | 
apply (erule cd_diff_r)  | 
|
79  | 
done  | 
|
80  | 
||
81  | 
||
| 62042 | 82  | 
subsubsection \<open>pow\<close>  | 
| 19802 | 83  | 
|
84  | 
lemma sq_pow_div2 [simp]:  | 
|
85  | 
"m mod 2 = 0 ==> ((n::nat)*n)^(m div 2) = n^m"  | 
|
| 64246 | 86  | 
apply (simp add: power2_eq_square [symmetric] power_mult [symmetric] minus_mod_eq_mult_div [symmetric])  | 
| 19802 | 87  | 
done  | 
| 17278 | 88  | 
|
| 1335 | 89  | 
end  |