| author | haftmann | 
| Mon, 06 Jun 2016 21:28:46 +0200 | |
| changeset 63239 | d562c9948dee | 
| parent 63167 | 0909deb8059b | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Quotient_Examples/Quotient_Int.thy  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
2  | 
Author: Cezary Kaliszyk  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
Author: Christian Urban  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
4  | 
|
| 41467 | 5  | 
Integers based on Quotients, based on an older version by Larry  | 
6  | 
Paulson.  | 
|
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 41467 | 8  | 
|
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
9  | 
theory Quotient_Int  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40928 
diff
changeset
 | 
10  | 
imports "~~/src/HOL/Library/Quotient_Product" Nat  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
11  | 
begin  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
13  | 
fun  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
14  | 
intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
15  | 
where  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
16  | 
"intrel (x, y) (u, v) = (x + v = u + y)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
18  | 
quotient_type int = "nat \<times> nat" / intrel  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
19  | 
by (auto simp add: equivp_def fun_eq_iff)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
21  | 
instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}"
 | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
22  | 
begin  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
24  | 
quotient_definition  | 
| 61076 | 25  | 
"0 :: int" is "(0::nat, 0::nat)" done  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
27  | 
quotient_definition  | 
| 61076 | 28  | 
"1 :: int" is "(1::nat, 0::nat)" done  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
30  | 
fun  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
31  | 
plus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
32  | 
where  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
33  | 
"plus_int_raw (x, y) (u, v) = (x + u, y + v)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
quotient_definition  | 
| 61076 | 36  | 
"(op +) :: (int \<Rightarrow> int \<Rightarrow> int)" is "plus_int_raw" by auto  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
38  | 
fun  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
39  | 
uminus_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
40  | 
where  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
41  | 
"uminus_int_raw (x, y) = (y, x)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
43  | 
quotient_definition  | 
| 61076 | 44  | 
"(uminus :: (int \<Rightarrow> int))" is "uminus_int_raw" by auto  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
46  | 
definition  | 
| 61076 | 47  | 
minus_int_def: "z - w = z + (-w::int)"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
49  | 
fun  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
50  | 
times_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
51  | 
where  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
52  | 
"times_int_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
53  | 
|
| 47092 | 54  | 
lemma times_int_raw_fst:  | 
55  | 
assumes a: "x \<approx> z"  | 
|
56  | 
shows "times_int_raw x y \<approx> times_int_raw z y"  | 
|
57  | 
using a  | 
|
58  | 
apply(cases x, cases y, cases z)  | 
|
59  | 
apply(auto simp add: times_int_raw.simps intrel.simps)  | 
|
| 
57492
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
54863 
diff
changeset
 | 
60  | 
apply(hypsubst_thin)  | 
| 47092 | 61  | 
apply(rename_tac u v w x y z)  | 
62  | 
apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x")  | 
|
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
63  | 
apply(simp add: ac_simps)  | 
| 47092 | 64  | 
apply(simp add: add_mult_distrib [symmetric])  | 
65  | 
done  | 
|
66  | 
||
67  | 
lemma times_int_raw_snd:  | 
|
68  | 
assumes a: "x \<approx> z"  | 
|
69  | 
shows "times_int_raw y x \<approx> times_int_raw y z"  | 
|
70  | 
using a  | 
|
71  | 
apply(cases x, cases y, cases z)  | 
|
72  | 
apply(auto simp add: times_int_raw.simps intrel.simps)  | 
|
| 
57492
 
74bf65a1910a
Hypsubst preserves equality hypotheses
 
Thomas Sewell <thomas.sewell@nicta.com.au> 
parents: 
54863 
diff
changeset
 | 
73  | 
apply(hypsubst_thin)  | 
| 47092 | 74  | 
apply(rename_tac u v w x y z)  | 
75  | 
apply(subgoal_tac "u*w + z*w = y*w + v*w & u*x + z*x = y*x + v*x")  | 
|
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
76  | 
apply(simp add: ac_simps)  | 
| 47092 | 77  | 
apply(simp add: add_mult_distrib [symmetric])  | 
78  | 
done  | 
|
79  | 
||
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
80  | 
quotient_definition  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
81  | 
"(op *) :: (int \<Rightarrow> int \<Rightarrow> int)" is "times_int_raw"  | 
| 47092 | 82  | 
apply(rule equivp_transp[OF int_equivp])  | 
83  | 
apply(rule times_int_raw_fst)  | 
|
84  | 
apply(assumption)  | 
|
85  | 
apply(rule times_int_raw_snd)  | 
|
86  | 
apply(assumption)  | 
|
87  | 
done  | 
|
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
88  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
89  | 
fun  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
90  | 
le_int_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
91  | 
where  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
92  | 
"le_int_raw (x, y) (u, v) = (x+v \<le> u+y)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
94  | 
quotient_definition  | 
| 47092 | 95  | 
le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" is "le_int_raw" by auto  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
97  | 
definition  | 
| 61076 | 98  | 
less_int_def: "(z::int) < w = (z \<le> w \<and> z \<noteq> w)"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
100  | 
definition  | 
| 61076 | 101  | 
zabs_def: "\<bar>i::int\<bar> = (if i < 0 then - i else i)"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
103  | 
definition  | 
| 61076 | 104  | 
zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
106  | 
instance ..  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
108  | 
end  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
110  | 
|
| 63167 | 111  | 
text\<open>The integers form a \<open>comm_ring_1\<close>\<close>  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
112  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
113  | 
instance int :: comm_ring_1  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
114  | 
proof  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
115  | 
fix i j k :: int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
116  | 
show "(i + j) + k = i + (j + k)"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
117  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
118  | 
show "i + j = j + i"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
119  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
120  | 
show "0 + i = (i::int)"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
121  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
122  | 
show "- i + i = 0"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
123  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
124  | 
show "i - j = i + - j"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
125  | 
by (simp add: minus_int_def)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
126  | 
show "(i * j) * k = i * (j * k)"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
127  | 
by (descending) (auto simp add: algebra_simps)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
128  | 
show "i * j = j * i"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
129  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
130  | 
show "1 * i = i"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
131  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
132  | 
show "(i + j) * k = i * k + j * k"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
133  | 
by (descending) (auto simp add: algebra_simps)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
134  | 
show "0 \<noteq> (1::int)"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
135  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
136  | 
qed  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
138  | 
lemma plus_int_raw_rsp_aux:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
139  | 
assumes a: "a \<approx> b" "c \<approx> d"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
140  | 
shows "plus_int_raw a c \<approx> plus_int_raw b d"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
141  | 
using a  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
142  | 
by (cases a, cases b, cases c, cases d)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
143  | 
(simp)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
144  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
145  | 
lemma add_abs_int:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
146  | 
"(abs_int (x,y)) + (abs_int (u,v)) =  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
147  | 
(abs_int (x + u, y + v))"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
148  | 
apply(simp add: plus_int_def id_simps)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
149  | 
apply(fold plus_int_raw.simps)  | 
| 47308 | 150  | 
apply(rule Quotient3_rel_abs[OF Quotient3_int])  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
151  | 
apply(rule plus_int_raw_rsp_aux)  | 
| 47308 | 152  | 
apply(simp_all add: rep_abs_rsp_left[OF Quotient3_int])  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
153  | 
done  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
155  | 
definition int_of_nat_raw:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
156  | 
"int_of_nat_raw m = (m :: nat, 0 :: nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
158  | 
quotient_definition  | 
| 47092 | 159  | 
"int_of_nat :: nat \<Rightarrow> int" is "int_of_nat_raw" done  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
161  | 
lemma int_of_nat:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
162  | 
shows "of_nat m = int_of_nat m"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
163  | 
by (induct m)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
164  | 
(simp_all add: zero_int_def one_int_def int_of_nat_def int_of_nat_raw add_abs_int)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
166  | 
instance int :: linorder  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
167  | 
proof  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
168  | 
fix i j k :: int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
169  | 
show antisym: "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
170  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
171  | 
show "(i < j) = (i \<le> j \<and> \<not> j \<le> i)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
172  | 
by (auto simp add: less_int_def dest: antisym)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
173  | 
show "i \<le> i"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
174  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
175  | 
show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
176  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
177  | 
show "i \<le> j \<or> j \<le> i"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
178  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
179  | 
qed  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
180  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
181  | 
instantiation int :: distrib_lattice  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
182  | 
begin  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
184  | 
definition  | 
| 61076 | 185  | 
"(inf :: int \<Rightarrow> int \<Rightarrow> int) = min"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
186  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
187  | 
definition  | 
| 61076 | 188  | 
"(sup :: int \<Rightarrow> int \<Rightarrow> int) = max"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
190  | 
instance  | 
| 61169 | 191  | 
by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
192  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
193  | 
end  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
194  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
195  | 
instance int :: ordered_cancel_ab_semigroup_add  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
196  | 
proof  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
197  | 
fix i j k :: int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
198  | 
show "i \<le> j \<Longrightarrow> k + i \<le> k + j"  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
199  | 
by (descending) (auto)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
200  | 
qed  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
202  | 
abbreviation  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
203  | 
"less_int_raw i j \<equiv> le_int_raw i j \<and> \<not>(i \<approx> j)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
204  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
205  | 
lemma zmult_zless_mono2_lemma:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
206  | 
fixes i j::int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
207  | 
and k::nat  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
208  | 
shows "i < j \<Longrightarrow> 0 < k \<Longrightarrow> of_nat k * i < of_nat k * j"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
209  | 
apply(induct "k")  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
210  | 
apply(simp)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
211  | 
apply(case_tac "k = 0")  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
47308 
diff
changeset
 | 
212  | 
apply(simp_all add: distrib_right add_strict_mono)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
213  | 
done  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
215  | 
lemma zero_le_imp_eq_int_raw:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
216  | 
fixes k::"(nat \<times> nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
217  | 
shows "less_int_raw (0, 0) k \<Longrightarrow> (\<exists>n > 0. k \<approx> int_of_nat_raw n)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
218  | 
apply(cases k)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
219  | 
apply(simp add:int_of_nat_raw)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
220  | 
apply(auto)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
221  | 
apply(rule_tac i="b" and j="a" in less_Suc_induct)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
222  | 
apply(auto)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
223  | 
done  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
225  | 
lemma zero_le_imp_eq_int:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
226  | 
fixes k::int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
227  | 
shows "0 < k \<Longrightarrow> \<exists>n > 0. k = of_nat n"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
228  | 
unfolding less_int_def int_of_nat  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
229  | 
by (descending) (rule zero_le_imp_eq_int_raw)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
231  | 
lemma zmult_zless_mono2:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
232  | 
fixes i j k::int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
233  | 
assumes a: "i < j" "0 < k"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
234  | 
shows "k * i < k * j"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
235  | 
using a  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
236  | 
by (drule_tac zero_le_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
237  | 
|
| 63167 | 238  | 
text\<open>The integers form an ordered integral domain\<close>  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
239  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
240  | 
instance int :: linordered_idom  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
241  | 
proof  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
242  | 
fix i j k :: int  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
243  | 
show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
244  | 
by (rule zmult_zless_mono2)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
245  | 
show "\<bar>i\<bar> = (if i < 0 then -i else i)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
246  | 
by (simp only: zabs_def)  | 
| 61076 | 247  | 
show "sgn (i::int) = (if i=0 then 0 else if 0<i then 1 else - 1)"  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
248  | 
by (simp only: zsgn_def)  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
249  | 
qed  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
250  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
251  | 
lemmas int_distrib =  | 
| 
49962
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
47308 
diff
changeset
 | 
252  | 
distrib_right [of z1 z2 w]  | 
| 
 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 
webertj 
parents: 
47308 
diff
changeset
 | 
253  | 
distrib_left [of w z1 z2]  | 
| 45605 | 254  | 
left_diff_distrib [of z1 z2 w]  | 
255  | 
right_diff_distrib [of w z1 z2]  | 
|
256  | 
minus_add_distrib[of z1 z2]  | 
|
257  | 
for z1 z2 w :: int  | 
|
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
258  | 
|
| 47304 | 259  | 
lemma int_induct2:  | 
260  | 
assumes "P 0 0"  | 
|
261  | 
and "\<And>n m. P n m \<Longrightarrow> P (Suc n) m"  | 
|
262  | 
and "\<And>n m. P n m \<Longrightarrow> P n (Suc m)"  | 
|
263  | 
shows "P n m"  | 
|
264  | 
using assms  | 
|
265  | 
by (induction_schema) (pat_completeness, lexicographic_order)  | 
|
266  | 
||
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
267  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
268  | 
lemma int_induct:  | 
| 47304 | 269  | 
fixes j :: int  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
270  | 
assumes a: "P 0"  | 
| 47304 | 271  | 
and b: "\<And>i::int. P i \<Longrightarrow> P (i + 1)"  | 
272  | 
and c: "\<And>i::int. P i \<Longrightarrow> P (i - 1)"  | 
|
273  | 
shows "P j"  | 
|
274  | 
using a b c  | 
|
275  | 
unfolding minus_int_def  | 
|
276  | 
by (descending) (auto intro: int_induct2)  | 
|
277  | 
||
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
278  | 
|
| 63167 | 279  | 
text \<open>Magnitide of an Integer, as a Natural Number: @{term nat}\<close>
 | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
280  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
281  | 
definition  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
282  | 
"int_to_nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
284  | 
quotient_definition  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
285  | 
"int_to_nat::int \<Rightarrow> nat"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
286  | 
is  | 
| 47304 | 287  | 
"int_to_nat_raw"  | 
288  | 
unfolding int_to_nat_raw_def by auto  | 
|
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
289  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
290  | 
lemma nat_le_eq_zle:  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
291  | 
fixes w z::"int"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
292  | 
shows "0 < w \<or> 0 \<le> z \<Longrightarrow> (int_to_nat w \<le> int_to_nat z) = (w \<le> z)"  | 
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
293  | 
unfolding less_int_def  | 
| 
37594
 
32ad67684ee7
cosmetics: avoided statement of raw theorems, used the method descending instead
 
Christian Urban <urbanc@in.tum.de> 
parents: 
36524 
diff
changeset
 | 
294  | 
by (descending) (auto simp add: int_to_nat_raw_def)  | 
| 
36524
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
295  | 
|
| 
 
3909002beca5
Tuning the quotient examples
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
296  | 
end  |