| author | wenzelm | 
| Thu, 01 Sep 2016 11:25:48 +0200 | |
| changeset 63742 | 1e676fcd7ede | 
| parent 62390 | 842917225d56 | 
| child 64245 | 3d00821444fc | 
| permissions | -rw-r--r-- | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Word/Word_Miscellaneous.thy  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
2  | 
Author: Miscellaneous  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
4  | 
|
| 61799 | 5  | 
section \<open>Miscellaneous lemmas, of at least doubtful value\<close>  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
theory Word_Miscellaneous  | 
| 58770 | 8  | 
imports Main "~~/src/HOL/Library/Bit" Misc_Numeric  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
11  | 
lemma power_minus_simp:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
12  | 
"0 < n \<Longrightarrow> a ^ n = a * a ^ (n - 1)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
by (auto dest: gr0_implies_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
15  | 
lemma funpow_minus_simp:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
"0 < n \<Longrightarrow> f ^^ n = f \<circ> f ^^ (n - 1)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
by (auto dest: gr0_implies_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
lemma power_numeral:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
20  | 
"a ^ numeral k = a * a ^ (pred_numeral k)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
21  | 
by (simp add: numeral_eq_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
23  | 
lemma funpow_numeral [simp]:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
24  | 
"f ^^ numeral k = f \<circ> f ^^ (pred_numeral k)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
25  | 
by (simp add: numeral_eq_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
27  | 
lemma replicate_numeral [simp]:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
28  | 
"replicate (numeral k) x = x # replicate (pred_numeral k) x"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
by (simp add: numeral_eq_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
31  | 
lemma rco_alt: "(f o g) ^^ n o f = f o (g o f) ^^ n"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
32  | 
apply (rule ext)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
33  | 
apply (induct n)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
34  | 
apply (simp_all add: o_def)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
35  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
37  | 
lemma list_exhaust_size_gt0:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
assumes y: "\<And>a list. y = a # list \<Longrightarrow> P"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
39  | 
shows "0 < length y \<Longrightarrow> P"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
40  | 
apply (cases y, simp)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
41  | 
apply (rule y)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
42  | 
apply fastforce  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
43  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
45  | 
lemma list_exhaust_size_eq0:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
46  | 
assumes y: "y = [] \<Longrightarrow> P"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
47  | 
shows "length y = 0 \<Longrightarrow> P"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
48  | 
apply (cases y)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
apply (rule y, simp)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
50  | 
apply simp  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
51  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
53  | 
lemma size_Cons_lem_eq:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
54  | 
"y = xa # list ==> size y = Suc k ==> size list = k"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
55  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
56  | 
|
| 62390 | 57  | 
lemmas ls_splits = prod.split prod.split_asm if_split_asm  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
59  | 
lemma not_B1_is_B0: "y \<noteq> (1::bit) \<Longrightarrow> y = (0::bit)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
60  | 
by (cases y) auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
62  | 
lemma B1_ass_B0:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
63  | 
assumes y: "y = (0::bit) \<Longrightarrow> y = (1::bit)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
64  | 
shows "y = (1::bit)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
65  | 
apply (rule classical)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
66  | 
apply (drule not_B1_is_B0)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
67  | 
apply (erule y)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
68  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
69  | 
|
| 61799 | 70  | 
\<comment> "simplifications for specific word lengths"  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
71  | 
lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc'  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
73  | 
lemmas s2n_ths = n2s_ths [symmetric]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
74  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
75  | 
lemma and_len: "xs = ys ==> xs = ys & length xs = length ys"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
76  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
78  | 
lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
79  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
80  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
81  | 
lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
82  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
84  | 
lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
85  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
87  | 
lemma if_Not_x: "(if p then ~ x else x) = (p = (~ x))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
88  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
90  | 
lemma if_x_Not: "(if p then x else ~ x) = (p = x)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
91  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
93  | 
lemma if_same_and: "(If p x y & If p u v) = (if p then x & u else y & v)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
94  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
96  | 
lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = (u) else y = (v))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
97  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
99  | 
lemma if_same_eq_not:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
100  | 
"(If p x y = (~ If p u v)) = (if p then x = (~u) else y = (~v))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
101  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
103  | 
(* note - if_Cons can cause blowup in the size, if p is complex,  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
104  | 
so make a simproc *)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
105  | 
lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
106  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
108  | 
lemma if_single:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
109  | 
"(if xc then [xab] else [an]) = [if xc then xab else an]"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
110  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
112  | 
lemma if_bool_simps:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
113  | 
"If p True y = (p | y) & If p False y = (~p & y) &  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
114  | 
If p y True = (p --> y) & If p y False = (p & y)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
115  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
116  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
117  | 
lemmas if_simps = if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
119  | 
lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
121  | 
lemma the_elemI: "y = {x} ==> the_elem y = x" 
 | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
122  | 
by simp  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
124  | 
lemma nonemptyE: "S ~= {} ==> (!!x. x : S ==> R) ==> R" by auto
 | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
125  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
126  | 
lemma gt_or_eq_0: "0 < y \<or> 0 = (y::nat)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
127  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
128  | 
lemmas xtr1 = xtrans(1)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
129  | 
lemmas xtr2 = xtrans(2)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
130  | 
lemmas xtr3 = xtrans(3)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
131  | 
lemmas xtr4 = xtrans(4)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
132  | 
lemmas xtr5 = xtrans(5)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
133  | 
lemmas xtr6 = xtrans(6)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
134  | 
lemmas xtr7 = xtrans(7)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
135  | 
lemmas xtr8 = xtrans(8)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
137  | 
lemmas nat_simps = diff_add_inverse2 diff_add_inverse  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
138  | 
lemmas nat_iffs = le_add1 le_add2  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
140  | 
lemma sum_imp_diff: "j = k + i ==> j - i = (k :: nat)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
141  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
142  | 
lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
143  | 
lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
144  | 
|
| 58681 | 145  | 
lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1"  | 
146  | 
by arith  | 
|
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
147  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
148  | 
lemmas eme1p = emep1 [simplified add.commute]  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
150  | 
lemma le_diff_eq': "(a \<le> c - b) = (b + a \<le> (c::int))" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
152  | 
lemma less_diff_eq': "(a < c - b) = (b + a < (c::int))" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
154  | 
lemma diff_less_eq': "(a - b < c) = (a < b + (c::int))" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
155  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
156  | 
lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
158  | 
lemma z1pdiv2:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
159  | 
"(2 * b + 1) div 2 = (b::int)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
161  | 
lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2,  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
162  | 
simplified int_one_le_iff_zero_less, simplified]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
164  | 
lemma axxbyy:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
165  | 
"a + m + m = b + n + n ==> (a = 0 | a = 1) ==> (b = 0 | b = 1) ==>  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
166  | 
a = b & m = (n :: int)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
168  | 
lemma axxmod2:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
169  | 
"(1 + x + x) mod 2 = (1 :: int) & (0 + x + x) mod 2 = (0 :: int)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
170  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
171  | 
lemma axxdiv2:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
172  | 
"(1 + x + x) div 2 = (x :: int) & (0 + x + x) div 2 = (x :: int)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
174  | 
lemmas iszero_minus = trans [THEN trans,  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
175  | 
OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
176  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
177  | 
lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add.commute]  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
178  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
179  | 
lemmas add_diff_cancel2 = add.commute [THEN diff_eq_eq [THEN iffD2]]  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
180  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
181  | 
lemmas rdmods [symmetric] = mod_minus_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
182  | 
mod_diff_left_eq mod_diff_right_eq mod_add_left_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
183  | 
mod_add_right_eq mod_mult_right_eq mod_mult_left_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
185  | 
lemma mod_plus_right:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
186  | 
"((a + x) mod m = (b + x) mod m) = (a mod m = b mod (m :: nat))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
187  | 
apply (induct x)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
188  | 
apply (simp_all add: mod_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
189  | 
apply arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
190  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
192  | 
lemma nat_minus_mod: "(n - n mod m) mod m = (0 :: nat)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
193  | 
by (induct n) (simp_all add : mod_Suc)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
194  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
195  | 
lemmas nat_minus_mod_plus_right = trans [OF nat_minus_mod mod_0 [symmetric],  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
196  | 
THEN mod_plus_right [THEN iffD2], simplified]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
197  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
198  | 
lemmas push_mods' = mod_add_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
199  | 
mod_mult_eq mod_diff_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
200  | 
mod_minus_eq  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
202  | 
lemmas push_mods = push_mods' [THEN eq_reflection]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
203  | 
lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
204  | 
lemmas mod_simps =  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
205  | 
mod_mult_self2_is_0 [THEN eq_reflection]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
206  | 
mod_mult_self1_is_0 [THEN eq_reflection]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
207  | 
mod_mod_trivial [THEN eq_reflection]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
208  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
209  | 
lemma nat_mod_eq:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
210  | 
"!!b. b < n ==> a mod n = b mod n ==> a mod n = (b :: nat)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
211  | 
by (induct a) auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
212  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
213  | 
lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
215  | 
lemma nat_mod_lem:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
216  | 
"(0 :: nat) < n ==> b < n = (b mod n = b)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
217  | 
apply safe  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
218  | 
apply (erule nat_mod_eq')  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
219  | 
apply (erule subst)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
220  | 
apply (erule mod_less_divisor)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
221  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
222  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
223  | 
lemma mod_nat_add:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
224  | 
"(x :: nat) < z ==> y < z ==>  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
225  | 
(x + y) mod z = (if x + y < z then x + y else x + y - z)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
226  | 
apply (rule nat_mod_eq)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
227  | 
apply auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
228  | 
apply (rule trans)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
229  | 
apply (rule le_mod_geq)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
230  | 
apply simp  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
231  | 
apply (rule nat_mod_eq')  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
232  | 
apply arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
233  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
234  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
235  | 
lemma mod_nat_sub:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
236  | 
"(x :: nat) < z ==> (x - y) mod z = x - y"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
237  | 
by (rule nat_mod_eq') arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
239  | 
lemma int_mod_eq:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
240  | 
"(0 :: int) <= b ==> b < n ==> a mod n = b mod n ==> a mod n = b"  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55417 
diff
changeset
 | 
241  | 
by (metis mod_pos_pos_trivial)  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
242  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55417 
diff
changeset
 | 
243  | 
lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *)  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
245  | 
lemma int_mod_le: "(0::int) <= a ==> a mod n <= a"  | 
| 58681 | 246  | 
by (fact Divides.semiring_numeral_div_class.mod_less_eq_dividend) (* FIXME: delete *)  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
248  | 
lemma mod_add_if_z:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
249  | 
"(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==>  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
250  | 
(x + y) mod z = (if x + y < z then x + y else x + y - z)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
251  | 
by (auto intro: int_mod_eq)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
252  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
253  | 
lemma mod_sub_if_z:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
254  | 
"(x :: int) < z ==> y < z ==> 0 <= y ==> 0 <= x ==> 0 <= z ==>  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
255  | 
(x - y) mod z = (if y <= x then x - y else x - y + z)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
256  | 
by (auto intro: int_mod_eq)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
258  | 
lemmas zmde = zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2], symmetric]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
259  | 
lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
261  | 
(* already have this for naturals, div_mult_self1/2, but not for ints *)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
262  | 
lemma zdiv_mult_self: "m ~= (0 :: int) ==> (a + m * n) div m = a div m + n"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
263  | 
apply (rule mcl)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
264  | 
prefer 2  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
265  | 
apply (erule asm_rl)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
266  | 
apply (simp add: zmde ring_distribs)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
267  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
268  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
269  | 
lemma mod_power_lem:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
270  | 
"a > 1 ==> a ^ n mod a ^ m = (if m <= n then 0 else (a :: int) ^ n)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
271  | 
apply clarsimp  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
272  | 
apply safe  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
273  | 
apply (simp add: dvd_eq_mod_eq_0 [symmetric])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
274  | 
apply (drule le_iff_add [THEN iffD1])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
275  | 
apply (force simp: power_add)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
276  | 
apply (rule mod_pos_pos_trivial)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
277  | 
apply (simp)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
278  | 
apply (rule power_strict_increasing)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
279  | 
apply auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
280  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
281  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
282  | 
lemma pl_pl_rels:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
283  | 
"a + b = c + d ==>  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
284  | 
a >= c & b <= d | a <= c & b >= (d :: nat)" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
285  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
286  | 
lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels]  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
287  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
288  | 
lemma minus_eq: "(m - k = m) = (k = 0 | m = (0 :: nat))" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
289  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
290  | 
lemma pl_pl_mm: "(a :: nat) + b = c + d ==> a - c = d - b" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
291  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
292  | 
lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm]  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
293  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
294  | 
lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
295  | 
lemmas dtle = xtr3 [OF dme [symmetric] le_add1]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
296  | 
lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] dtle]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
297  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
298  | 
lemma td_gal:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
299  | 
"0 < c ==> (a >= b * c) = (a div c >= (b :: nat))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
300  | 
apply safe  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
301  | 
apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
302  | 
apply (erule th2)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
303  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
304  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
305  | 
lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
306  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
307  | 
lemma div_mult_le: "(a :: nat) div b * b <= a"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
308  | 
by (fact dtle)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
309  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
310  | 
lemmas sdl = split_div_lemma [THEN iffD1, symmetric]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
312  | 
lemma given_quot: "f > (0 :: nat) ==> (f * l + (f - 1)) div f = l"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
313  | 
by (rule sdl, assumption) (simp (no_asm))  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
314  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
315  | 
lemma given_quot_alt: "f > (0 :: nat) ==> (l * f + f - Suc 0) div f = l"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
316  | 
apply (frule given_quot)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
317  | 
apply (rule trans)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
318  | 
prefer 2  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
319  | 
apply (erule asm_rl)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
320  | 
apply (rule_tac f="%n. n div f" in arg_cong)  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
321  | 
apply (simp add : ac_simps)  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
322  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
323  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
324  | 
lemma diff_mod_le: "(a::nat) < d ==> b dvd d ==> a - a mod b <= d - b"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
325  | 
apply (unfold dvd_def)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
326  | 
apply clarify  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
327  | 
apply (case_tac k)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
328  | 
apply clarsimp  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
329  | 
apply clarify  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
330  | 
apply (cases "b > 0")  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
55816 
diff
changeset
 | 
331  | 
apply (drule mult.commute [THEN xtr1])  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
332  | 
apply (frule (1) td_gal_lt [THEN iffD1])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
333  | 
apply (clarsimp simp: le_simps)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
334  | 
apply (rule mult_div_cancel [THEN [2] xtr4])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
335  | 
apply (rule mult_mono)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
336  | 
apply auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
337  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
338  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
339  | 
lemma less_le_mult':  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
340  | 
"w * c < b * c ==> 0 \<le> c ==> (w + 1) * c \<le> b * (c::int)"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
341  | 
apply (rule mult_right_mono)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
342  | 
apply (rule zless_imp_add1_zle)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
343  | 
apply (erule (1) mult_right_less_imp_less)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
344  | 
apply assumption  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
345  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
346  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55417 
diff
changeset
 | 
347  | 
lemma less_le_mult:  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55417 
diff
changeset
 | 
348  | 
"w * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> w * c + c \<le> b * (c::int)"  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55417 
diff
changeset
 | 
349  | 
using less_le_mult' [of w c b] by (simp add: algebra_simps)  | 
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
350  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
351  | 
lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult,  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
352  | 
simplified left_diff_distrib]  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
353  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
354  | 
lemma gen_minus: "0 < n ==> f n = f (Suc (n - 1))"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
355  | 
by auto  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
356  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
357  | 
lemma mpl_lem: "j <= (i :: nat) ==> k < j ==> i - j + k < i" by arith  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
358  | 
|
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
359  | 
lemma nonneg_mod_div:  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
360  | 
"0 <= a ==> 0 <= b ==> 0 <= (a mod b :: int) & 0 <= a div b"  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
361  | 
apply (cases "b = 0", clarsimp)  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
362  | 
apply (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2])  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
363  | 
done  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
364  | 
|
| 54872 | 365  | 
declare iszero_0 [intro]  | 
366  | 
||
367  | 
lemma min_pm [simp]:  | 
|
368  | 
"min a b + (a - b) = (a :: nat)"  | 
|
369  | 
by arith  | 
|
370  | 
||
371  | 
lemma min_pm1 [simp]:  | 
|
372  | 
"a - b + min a b = (a :: nat)"  | 
|
373  | 
by arith  | 
|
374  | 
||
375  | 
lemma rev_min_pm [simp]:  | 
|
376  | 
"min b a + (a - b) = (a :: nat)"  | 
|
377  | 
by arith  | 
|
378  | 
||
379  | 
lemma rev_min_pm1 [simp]:  | 
|
380  | 
"a - b + min b a = (a :: nat)"  | 
|
381  | 
by arith  | 
|
382  | 
||
383  | 
lemma min_minus [simp]:  | 
|
384  | 
"min m (m - k) = (m - k :: nat)"  | 
|
385  | 
by arith  | 
|
386  | 
||
387  | 
lemma min_minus' [simp]:  | 
|
388  | 
"min (m - k) m = (m - k :: nat)"  | 
|
389  | 
by arith  | 
|
390  | 
||
| 
53062
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
391  | 
end  | 
| 
 
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
 
haftmann 
parents:  
diff
changeset
 | 
392  |