author | wenzelm |
Tue, 17 Sep 2013 17:17:55 +0200 | |
changeset 53685 | 983711bc98e0 |
parent 53062 | 3af1a6020014 |
child 54871 | 51612b889361 |
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 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
5 |
header {* Miscellaneous lemmas, of at least doubtful value *} |
3af1a6020014
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 |
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
8 |
imports Main Parity "~~/src/HOL/Library/Bit" Misc_Numeric |
3af1a6020014
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 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
57 |
lemmas ls_splits = prod.split prod.split_asm split_if_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
|
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 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
70 |
-- "simplifications for specific word lengths" |
3af1a6020014
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 |
(* TODO: move name bindings to List.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
|
122 |
lemmas tl_Nil = tl.simps (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
|
123 |
lemmas tl_Cons = tl.simps (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
|
124 |
|
3af1a6020014
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 |
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
|
126 |
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
|
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 |
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
|
129 |
|
3af1a6020014
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 |
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
|
131 |
|
3af1a6020014
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 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
|
133 |
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
|
134 |
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
|
135 |
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
|
136 |
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
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
|
3af1a6020014
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 |
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
|
142 |
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
|
143 |
|
3af1a6020014
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 |
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
|
145 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
146 |
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
|
147 |
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
|
148 |
|
3af1a6020014
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 |
lemma nmod2: "n mod (2::int) = 0 | n mod 2 = 1" 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
|
150 |
|
3af1a6020014
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 |
lemma emep1: |
3af1a6020014
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 |
"even n ==> even d ==> 0 <= d ==> (n + 1) mod (d :: int) = (n mod d) + 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
|
153 |
apply (simp add: add_commute) |
3af1a6020014
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 |
apply (safe dest!: even_equiv_def [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
|
155 |
apply (subst pos_zmod_mult_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
|
156 |
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
|
157 |
apply (simp add: mod_mult_mult1) |
3af1a6020014
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 |
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
|
159 |
|
3af1a6020014
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 |
lemmas eme1p = emep1 [simplified add_commute] |
3af1a6020014
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 |
|
3af1a6020014
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 |
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
|
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 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
|
165 |
|
3af1a6020014
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 |
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
|
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 |
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
|
169 |
lemmas z1pmod2' = zero_le_one [THEN pos_zmod_mult_2, 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
|
170 |
lemmas z1pdiv2' = zero_le_one [THEN pos_zdiv_mult_2, 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
|
171 |
|
3af1a6020014
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 |
lemma z1pmod2: |
3af1a6020014
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 |
"(2 * b + 1) mod 2 = (1::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
|
174 |
|
3af1a6020014
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 |
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
|
176 |
"(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
|
177 |
|
3af1a6020014
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 |
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
|
179 |
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
|
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 |
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
|
182 |
"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
|
183 |
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
|
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 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
|
186 |
"(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
|
187 |
|
3af1a6020014
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 |
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
|
189 |
"(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
|
190 |
|
3af1a6020014
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 |
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
|
192 |
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
|
193 |
|
3af1a6020014
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 |
lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add_commute] |
3af1a6020014
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 |
|
3af1a6020014
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 |
lemmas add_diff_cancel2 = add_commute [THEN diff_eq_eq [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
|
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 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
|
199 |
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
|
200 |
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
|
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 |
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
|
203 |
"((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
|
204 |
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
|
205 |
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
|
206 |
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
|
207 |
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
|
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_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
|
210 |
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
|
211 |
|
3af1a6020014
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 |
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
|
213 |
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
|
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 |
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
|
216 |
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
|
217 |
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
|
218 |
|
3af1a6020014
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 |
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
|
220 |
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
|
221 |
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
|
222 |
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
|
223 |
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
|
224 |
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
|
225 |
|
3af1a6020014
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 |
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
|
227 |
"!!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
|
228 |
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
|
229 |
|
3af1a6020014
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 |
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
|
231 |
|
3af1a6020014
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 |
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
|
233 |
"(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
|
234 |
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
|
235 |
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
|
236 |
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
|
237 |
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
|
238 |
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
|
239 |
|
3af1a6020014
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 |
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
|
241 |
"(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
|
242 |
(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
|
243 |
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
|
244 |
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
|
245 |
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
|
246 |
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
|
247 |
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
|
248 |
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
|
249 |
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
|
250 |
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
|
251 |
|
3af1a6020014
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 |
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
|
253 |
"(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
|
254 |
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
|
255 |
|
3af1a6020014
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 |
lemma int_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
|
257 |
"(0 :: int) < n ==> (0 <= b & 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
|
258 |
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
|
259 |
apply (erule (1) 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
|
260 |
apply (erule_tac [!] 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
|
261 |
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
|
262 |
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
|
263 |
|
3af1a6020014
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 |
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
|
265 |
"(0 :: int) <= b ==> b < n ==> a mod n = b mod n ==> a 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
|
266 |
by clarsimp (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
|
267 |
|
3af1a6020014
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 |
lemmas int_mod_eq' = refl [THEN [3] 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
|
269 |
|
3af1a6020014
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 |
lemma int_mod_le: "(0::int) <= a ==> a mod n <= 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
|
271 |
by (fact zmod_le_nonneg_dividend) (* 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
|
272 |
|
3af1a6020014
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 |
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
|
274 |
"(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
|
275 |
(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
|
276 |
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
|
277 |
|
3af1a6020014
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 |
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
|
279 |
"(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
|
280 |
(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
|
281 |
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
|
282 |
|
3af1a6020014
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 |
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
|
284 |
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
|
285 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
286 |
(* 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
|
287 |
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
|
288 |
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
|
289 |
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
|
290 |
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
|
291 |
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
|
292 |
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
|
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 |
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
|
295 |
"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
|
296 |
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
|
297 |
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
|
298 |
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
|
299 |
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
|
300 |
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
|
301 |
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
|
302 |
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
|
303 |
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
|
304 |
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
|
305 |
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
|
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 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
|
308 |
"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
|
309 |
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
|
310 |
|
3af1a6020014
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 |
lemmas pl_pl_rels' = add_commute [THEN [2] trans, THEN 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
|
312 |
|
3af1a6020014
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 |
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
|
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 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
|
316 |
|
3af1a6020014
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 |
lemmas pl_pl_mm' = add_commute [THEN [2] trans, THEN pl_pl_mm] |
3af1a6020014
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 |
|
3af1a6020014
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 |
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
|
320 |
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
|
321 |
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
|
322 |
|
3af1a6020014
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 |
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
|
324 |
"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
|
325 |
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
|
326 |
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
|
327 |
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
|
328 |
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
|
329 |
|
3af1a6020014
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 |
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
|
331 |
|
3af1a6020014
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 |
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
|
333 |
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
|
334 |
|
3af1a6020014
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 |
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
|
336 |
|
3af1a6020014
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 |
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
|
338 |
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
|
339 |
|
3af1a6020014
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 |
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
|
341 |
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
|
342 |
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
|
343 |
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
|
344 |
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
|
345 |
apply (rule_tac f="%n. n div f" in arg_cong) |
3af1a6020014
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 |
apply (simp add : mult_ac) |
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
347 |
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
|
348 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
349 |
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
|
350 |
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
|
351 |
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
|
352 |
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
|
353 |
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
|
354 |
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
|
355 |
apply (cases "b > 0") |
3af1a6020014
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 |
apply (drule mult_commute [THEN xtr1]) |
3af1a6020014
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 |
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
|
358 |
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
|
359 |
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
|
360 |
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
|
361 |
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
|
362 |
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
|
363 |
|
3af1a6020014
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 |
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
|
365 |
"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
|
366 |
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
|
367 |
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
|
368 |
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
|
369 |
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
|
370 |
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
|
371 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
372 |
lemmas less_le_mult = less_le_mult' [simplified distrib_right, 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
|
373 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
374 |
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
|
375 |
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
|
376 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
377 |
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
|
378 |
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
|
379 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
380 |
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
|
381 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
382 |
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
|
383 |
"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
|
384 |
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
|
385 |
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
|
386 |
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
|
387 |
|
3af1a6020014
some vague grouping of related theorems, with slight tuning of headings and sorting out of dubious lemmas into separate theory
haftmann
parents:
diff
changeset
|
388 |
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
|
389 |