author | wenzelm |
Sat, 18 Aug 2007 13:32:28 +0200 | |
changeset 24325 | 5c29e8822f50 |
parent 24195 | 7d1a16c77f7c |
child 24506 | 020db6ec334a |
permissions | -rw-r--r-- |
17296 | 1 |
(* Title : HOL/Hyperreal/StarClasses.thy |
2 |
ID : $Id$ |
|
3 |
Author : Brian Huffman |
|
4 |
*) |
|
5 |
||
6 |
header {* Class Instances *} |
|
7 |
||
8 |
theory StarClasses |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
9 |
imports StarDef |
17296 | 10 |
begin |
11 |
||
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
12 |
subsection {* Syntactic classes *} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
13 |
|
22316 | 14 |
instance star :: (zero) zero |
15 |
star_zero_def: "0 \<equiv> star_of 0" .. |
|
16 |
||
17 |
instance star :: (one) one |
|
18 |
star_one_def: "1 \<equiv> star_of 1" .. |
|
19 |
||
20 |
instance star :: (plus) plus |
|
21 |
star_add_def: "(op +) \<equiv> *f2* (op +)" .. |
|
22 |
||
23 |
instance star :: (times) times |
|
24 |
star_mult_def: "(op *) \<equiv> *f2* (op *)" .. |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
25 |
|
22316 | 26 |
instance star :: (minus) minus |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
27 |
star_minus_def: "uminus \<equiv> *f* uminus" |
23879 | 28 |
star_diff_def: "(op -) \<equiv> *f2* (op -)" .. |
29 |
||
30 |
instance star :: (abs) abs |
|
22316 | 31 |
star_abs_def: "abs \<equiv> *f* abs" .. |
32 |
||
33 |
instance star :: (inverse) inverse |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
34 |
star_divide_def: "(op /) \<equiv> *f2* (op /)" |
22316 | 35 |
star_inverse_def: "inverse \<equiv> *f* inverse" .. |
36 |
||
37 |
instance star :: (number) number |
|
38 |
star_number_def: "number_of b \<equiv> star_of (number_of b)" .. |
|
39 |
||
22993 | 40 |
instance star :: (Divides.div) Divides.div |
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
41 |
star_div_def: "(op div) \<equiv> *f2* (op div)" |
22316 | 42 |
star_mod_def: "(op mod) \<equiv> *f2* (op mod)" .. |
43 |
||
44 |
instance star :: (power) power |
|
45 |
star_power_def: "(op ^) \<equiv> \<lambda>x n. ( *f* (\<lambda>x. x ^ n)) x" .. |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
46 |
|
22452
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
47 |
instance star :: (ord) ord |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
48 |
star_le_def: "(op \<le>) \<equiv> *p2* (op \<le>)" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
49 |
star_less_def: "(op <) \<equiv> *p2* (op <)" .. |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
50 |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
51 |
lemmas star_class_defs [transfer_unfold] = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
52 |
star_zero_def star_one_def star_number_def |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
53 |
star_add_def star_diff_def star_minus_def |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
54 |
star_mult_def star_divide_def star_inverse_def |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
55 |
star_le_def star_less_def star_abs_def |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
56 |
star_div_def star_mod_def star_power_def |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
57 |
|
20719 | 58 |
text {* Class operations preserve standard elements *} |
59 |
||
60 |
lemma Standard_zero: "0 \<in> Standard" |
|
61 |
by (simp add: star_zero_def) |
|
62 |
||
63 |
lemma Standard_one: "1 \<in> Standard" |
|
64 |
by (simp add: star_one_def) |
|
65 |
||
66 |
lemma Standard_number_of: "number_of b \<in> Standard" |
|
67 |
by (simp add: star_number_def) |
|
68 |
||
69 |
lemma Standard_add: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x + y \<in> Standard" |
|
70 |
by (simp add: star_add_def) |
|
71 |
||
72 |
lemma Standard_diff: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x - y \<in> Standard" |
|
73 |
by (simp add: star_diff_def) |
|
74 |
||
75 |
lemma Standard_minus: "x \<in> Standard \<Longrightarrow> - x \<in> Standard" |
|
76 |
by (simp add: star_minus_def) |
|
77 |
||
78 |
lemma Standard_mult: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x * y \<in> Standard" |
|
79 |
by (simp add: star_mult_def) |
|
80 |
||
81 |
lemma Standard_divide: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x / y \<in> Standard" |
|
82 |
by (simp add: star_divide_def) |
|
83 |
||
84 |
lemma Standard_inverse: "x \<in> Standard \<Longrightarrow> inverse x \<in> Standard" |
|
85 |
by (simp add: star_inverse_def) |
|
86 |
||
87 |
lemma Standard_abs: "x \<in> Standard \<Longrightarrow> abs x \<in> Standard" |
|
88 |
by (simp add: star_abs_def) |
|
89 |
||
90 |
lemma Standard_div: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x div y \<in> Standard" |
|
91 |
by (simp add: star_div_def) |
|
92 |
||
93 |
lemma Standard_mod: "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> x mod y \<in> Standard" |
|
94 |
by (simp add: star_mod_def) |
|
95 |
||
96 |
lemma Standard_power: "x \<in> Standard \<Longrightarrow> x ^ n \<in> Standard" |
|
97 |
by (simp add: star_power_def) |
|
98 |
||
99 |
lemmas Standard_simps [simp] = |
|
100 |
Standard_zero Standard_one Standard_number_of |
|
101 |
Standard_add Standard_diff Standard_minus |
|
102 |
Standard_mult Standard_divide Standard_inverse |
|
103 |
Standard_abs Standard_div Standard_mod |
|
104 |
Standard_power |
|
105 |
||
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
106 |
text {* @{term star_of} preserves class operations *} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
107 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
108 |
lemma star_of_add: "star_of (x + y) = star_of x + star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
109 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
110 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
111 |
lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
112 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
113 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
114 |
lemma star_of_minus: "star_of (-x) = - star_of x" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
115 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
116 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
117 |
lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
118 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
119 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
120 |
lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
121 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
122 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
123 |
lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
124 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
125 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
126 |
lemma star_of_div: "star_of (x div y) = star_of x div star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
127 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
128 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
129 |
lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
130 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
131 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
132 |
lemma star_of_power: "star_of (x ^ n) = star_of x ^ n" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
133 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
134 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
135 |
lemma star_of_abs: "star_of (abs x) = abs (star_of x)" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
136 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
137 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
138 |
text {* @{term star_of} preserves numerals *} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
139 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
140 |
lemma star_of_zero: "star_of 0 = 0" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
141 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
142 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
143 |
lemma star_of_one: "star_of 1 = 1" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
144 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
145 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
146 |
lemma star_of_number_of: "star_of (number_of x) = number_of x" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
147 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
148 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
149 |
text {* @{term star_of} preserves orderings *} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
150 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
151 |
lemma star_of_less: "(star_of x < star_of y) = (x < y)" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
152 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
153 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
154 |
lemma star_of_le: "(star_of x \<le> star_of y) = (x \<le> y)" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
155 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
156 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
157 |
lemma star_of_eq: "(star_of x = star_of y) = (x = y)" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
158 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
159 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
160 |
text{*As above, for 0*} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
161 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
162 |
lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
163 |
lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
164 |
lemmas star_of_0_eq = star_of_eq [of 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
165 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
166 |
lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
167 |
lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
168 |
lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
169 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
170 |
text{*As above, for 1*} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
171 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
172 |
lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
173 |
lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
174 |
lemmas star_of_1_eq = star_of_eq [of 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
175 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
176 |
lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
177 |
lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
178 |
lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
179 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
180 |
text{*As above, for numerals*} |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
181 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
182 |
lemmas star_of_number_less = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
183 |
star_of_less [of "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
184 |
lemmas star_of_number_le = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
185 |
star_of_le [of "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
186 |
lemmas star_of_number_eq = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
187 |
star_of_eq [of "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
188 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
189 |
lemmas star_of_less_number = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
190 |
star_of_less [of _ "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
191 |
lemmas star_of_le_number = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
192 |
star_of_le [of _ "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
193 |
lemmas star_of_eq_number = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
194 |
star_of_eq [of _ "number_of w", standard, simplified star_of_number_of] |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
195 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
196 |
lemmas star_of_simps [simp] = |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
197 |
star_of_add star_of_diff star_of_minus |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
198 |
star_of_mult star_of_divide star_of_inverse |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
199 |
star_of_div star_of_mod |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
200 |
star_of_power star_of_abs |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
201 |
star_of_zero star_of_one star_of_number_of |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
202 |
star_of_less star_of_le star_of_eq |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
203 |
star_of_0_less star_of_0_le star_of_0_eq |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
204 |
star_of_less_0 star_of_le_0 star_of_eq_0 |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
205 |
star_of_1_less star_of_1_le star_of_1_eq |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
206 |
star_of_less_1 star_of_le_1 star_of_eq_1 |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
207 |
star_of_number_less star_of_number_le star_of_number_eq |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
208 |
star_of_less_number star_of_le_number star_of_eq_number |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
209 |
|
22452
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
210 |
subsection {* Ordering and lattice classes *} |
17296 | 211 |
|
212 |
instance star :: (order) order |
|
213 |
apply (intro_classes) |
|
22316 | 214 |
apply (transfer, rule order_less_le) |
17296 | 215 |
apply (transfer, rule order_refl) |
216 |
apply (transfer, erule (1) order_trans) |
|
217 |
apply (transfer, erule (1) order_antisym) |
|
218 |
done |
|
219 |
||
22452
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
220 |
instance star :: (lower_semilattice) lower_semilattice |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
221 |
star_inf_def [transfer_unfold]: "inf \<equiv> *f2* inf" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
222 |
by default (transfer star_inf_def, auto)+ |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
223 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
224 |
instance star :: (upper_semilattice) upper_semilattice |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
225 |
star_sup_def [transfer_unfold]: "sup \<equiv> *f2* sup" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
226 |
by default (transfer star_sup_def, auto)+ |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
227 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
228 |
instance star :: (lattice) lattice .. |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
229 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
230 |
instance star :: (distrib_lattice) distrib_lattice |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
231 |
by default (transfer, auto simp add: sup_inf_distrib1) |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
232 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
233 |
lemma Standard_inf [simp]: |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
234 |
"\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> inf x y \<in> Standard" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
235 |
by (simp add: star_inf_def) |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
236 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
237 |
lemma Standard_sup [simp]: |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
238 |
"\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> sup x y \<in> Standard" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
239 |
by (simp add: star_sup_def) |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
240 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
241 |
lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
242 |
by transfer (rule refl) |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
243 |
|
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
244 |
lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)" |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
245 |
by transfer (rule refl) |
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents:
22422
diff
changeset
|
246 |
|
17296 | 247 |
instance star :: (linorder) linorder |
248 |
by (intro_classes, transfer, rule linorder_linear) |
|
249 |
||
20720 | 250 |
lemma star_max_def [transfer_unfold]: "max = *f2* max" |
251 |
apply (rule ext, rule ext) |
|
252 |
apply (unfold max_def, transfer, fold max_def) |
|
253 |
apply (rule refl) |
|
254 |
done |
|
255 |
||
256 |
lemma star_min_def [transfer_unfold]: "min = *f2* min" |
|
257 |
apply (rule ext, rule ext) |
|
258 |
apply (unfold min_def, transfer, fold min_def) |
|
259 |
apply (rule refl) |
|
260 |
done |
|
261 |
||
262 |
lemma Standard_max [simp]: |
|
263 |
"\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> max x y \<in> Standard" |
|
264 |
by (simp add: star_max_def) |
|
265 |
||
266 |
lemma Standard_min [simp]: |
|
267 |
"\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> min x y \<in> Standard" |
|
268 |
by (simp add: star_min_def) |
|
269 |
||
270 |
lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)" |
|
271 |
by transfer (rule refl) |
|
272 |
||
273 |
lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)" |
|
274 |
by transfer (rule refl) |
|
275 |
||
17296 | 276 |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
277 |
subsection {* Ordered group classes *} |
17296 | 278 |
|
279 |
instance star :: (semigroup_add) semigroup_add |
|
280 |
by (intro_classes, transfer, rule add_assoc) |
|
281 |
||
282 |
instance star :: (ab_semigroup_add) ab_semigroup_add |
|
283 |
by (intro_classes, transfer, rule add_commute) |
|
284 |
||
285 |
instance star :: (semigroup_mult) semigroup_mult |
|
286 |
by (intro_classes, transfer, rule mult_assoc) |
|
287 |
||
288 |
instance star :: (ab_semigroup_mult) ab_semigroup_mult |
|
289 |
by (intro_classes, transfer, rule mult_commute) |
|
290 |
||
291 |
instance star :: (comm_monoid_add) comm_monoid_add |
|
22384
33a46e6c7f04
prefix of class interpretation not mandatory any longer
haftmann
parents:
22316
diff
changeset
|
292 |
by (intro_classes, transfer, rule comm_monoid_add_class.zero_plus.add_0) |
17296 | 293 |
|
294 |
instance star :: (monoid_mult) monoid_mult |
|
295 |
apply (intro_classes) |
|
296 |
apply (transfer, rule mult_1_left) |
|
297 |
apply (transfer, rule mult_1_right) |
|
298 |
done |
|
299 |
||
300 |
instance star :: (comm_monoid_mult) comm_monoid_mult |
|
301 |
by (intro_classes, transfer, rule mult_1) |
|
302 |
||
303 |
instance star :: (cancel_semigroup_add) cancel_semigroup_add |
|
304 |
apply (intro_classes) |
|
305 |
apply (transfer, erule add_left_imp_eq) |
|
306 |
apply (transfer, erule add_right_imp_eq) |
|
307 |
done |
|
308 |
||
309 |
instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add |
|
310 |
by (intro_classes, transfer, rule add_imp_eq) |
|
311 |
||
312 |
instance star :: (ab_group_add) ab_group_add |
|
313 |
apply (intro_classes) |
|
314 |
apply (transfer, rule left_minus) |
|
315 |
apply (transfer, rule diff_minus) |
|
316 |
done |
|
317 |
||
318 |
instance star :: (pordered_ab_semigroup_add) pordered_ab_semigroup_add |
|
319 |
by (intro_classes, transfer, rule add_left_mono) |
|
320 |
||
321 |
instance star :: (pordered_cancel_ab_semigroup_add) pordered_cancel_ab_semigroup_add .. |
|
322 |
||
323 |
instance star :: (pordered_ab_semigroup_add_imp_le) pordered_ab_semigroup_add_imp_le |
|
324 |
by (intro_classes, transfer, rule add_le_imp_le_left) |
|
325 |
||
326 |
instance star :: (pordered_ab_group_add) pordered_ab_group_add .. |
|
327 |
instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. |
|
328 |
instance star :: (lordered_ab_group_meet) lordered_ab_group_meet .. |
|
329 |
instance star :: (lordered_ab_group_meet) lordered_ab_group_meet .. |
|
330 |
instance star :: (lordered_ab_group) lordered_ab_group .. |
|
331 |
||
332 |
instance star :: (lordered_ab_group_abs) lordered_ab_group_abs |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
333 |
by (intro_classes, transfer, rule abs_lattice) |
17296 | 334 |
|
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
335 |
subsection {* Ring and field classes *} |
17296 | 336 |
|
337 |
instance star :: (semiring) semiring |
|
338 |
apply (intro_classes) |
|
339 |
apply (transfer, rule left_distrib) |
|
340 |
apply (transfer, rule right_distrib) |
|
341 |
done |
|
342 |
||
21199
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents:
20720
diff
changeset
|
343 |
instance star :: (semiring_0) semiring_0 |
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents:
20720
diff
changeset
|
344 |
by intro_classes (transfer, simp)+ |
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents:
20720
diff
changeset
|
345 |
|
17296 | 346 |
instance star :: (semiring_0_cancel) semiring_0_cancel .. |
347 |
||
348 |
instance star :: (comm_semiring) comm_semiring |
|
349 |
by (intro_classes, transfer, rule distrib) |
|
350 |
||
351 |
instance star :: (comm_semiring_0) comm_semiring_0 .. |
|
352 |
instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. |
|
353 |
||
20633 | 354 |
instance star :: (zero_neq_one) zero_neq_one |
17296 | 355 |
by (intro_classes, transfer, rule zero_neq_one) |
356 |
||
357 |
instance star :: (semiring_1) semiring_1 .. |
|
358 |
instance star :: (comm_semiring_1) comm_semiring_1 .. |
|
359 |
||
20633 | 360 |
instance star :: (no_zero_divisors) no_zero_divisors |
17296 | 361 |
by (intro_classes, transfer, rule no_zero_divisors) |
362 |
||
363 |
instance star :: (semiring_1_cancel) semiring_1_cancel .. |
|
364 |
instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. |
|
365 |
instance star :: (ring) ring .. |
|
366 |
instance star :: (comm_ring) comm_ring .. |
|
367 |
instance star :: (ring_1) ring_1 .. |
|
368 |
instance star :: (comm_ring_1) comm_ring_1 .. |
|
22992 | 369 |
instance star :: (ring_no_zero_divisors) ring_no_zero_divisors .. |
23551
84f0996a530b
rename class dom to ring_1_no_zero_divisors (cf. HOL/Ring_and_Field.thy 1.84 by huffman);
wenzelm
parents:
23282
diff
changeset
|
370 |
instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. |
17296 | 371 |
instance star :: (idom) idom .. |
372 |
||
20540 | 373 |
instance star :: (division_ring) division_ring |
374 |
apply (intro_classes) |
|
375 |
apply (transfer, erule left_inverse) |
|
376 |
apply (transfer, erule right_inverse) |
|
377 |
done |
|
378 |
||
17296 | 379 |
instance star :: (field) field |
380 |
apply (intro_classes) |
|
381 |
apply (transfer, erule left_inverse) |
|
382 |
apply (transfer, rule divide_inverse) |
|
383 |
done |
|
384 |
||
385 |
instance star :: (division_by_zero) division_by_zero |
|
386 |
by (intro_classes, transfer, rule inverse_zero) |
|
387 |
||
388 |
instance star :: (pordered_semiring) pordered_semiring |
|
389 |
apply (intro_classes) |
|
390 |
apply (transfer, erule (1) mult_left_mono) |
|
391 |
apply (transfer, erule (1) mult_right_mono) |
|
392 |
done |
|
393 |
||
394 |
instance star :: (pordered_cancel_semiring) pordered_cancel_semiring .. |
|
395 |
||
396 |
instance star :: (ordered_semiring_strict) ordered_semiring_strict |
|
397 |
apply (intro_classes) |
|
398 |
apply (transfer, erule (1) mult_strict_left_mono) |
|
399 |
apply (transfer, erule (1) mult_strict_right_mono) |
|
400 |
done |
|
401 |
||
402 |
instance star :: (pordered_comm_semiring) pordered_comm_semiring |
|
23879 | 403 |
by (intro_classes, transfer, rule mult_mono1_class.less_eq_less_times_zero.mult_mono) |
17296 | 404 |
|
405 |
instance star :: (pordered_cancel_comm_semiring) pordered_cancel_comm_semiring .. |
|
406 |
||
407 |
instance star :: (ordered_comm_semiring_strict) ordered_comm_semiring_strict |
|
22518 | 408 |
by (intro_classes, transfer, rule ordered_comm_semiring_strict_class.plus_less_eq_less_zero_times.mult_strict_mono) |
17296 | 409 |
|
410 |
instance star :: (pordered_ring) pordered_ring .. |
|
411 |
instance star :: (lordered_ring) lordered_ring .. |
|
412 |
||
20633 | 413 |
instance star :: (abs_if) abs_if |
17296 | 414 |
by (intro_classes, transfer, rule abs_if) |
415 |
||
416 |
instance star :: (ordered_ring_strict) ordered_ring_strict .. |
|
417 |
instance star :: (pordered_comm_ring) pordered_comm_ring .. |
|
418 |
||
419 |
instance star :: (ordered_semidom) ordered_semidom |
|
420 |
by (intro_classes, transfer, rule zero_less_one) |
|
421 |
||
422 |
instance star :: (ordered_idom) ordered_idom .. |
|
423 |
instance star :: (ordered_field) ordered_field .. |
|
424 |
||
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
425 |
subsection {* Power classes *} |
17296 | 426 |
|
427 |
text {* |
|
428 |
Proving the class axiom @{thm [source] power_Suc} for type |
|
429 |
@{typ "'a star"} is a little tricky, because it quantifies |
|
430 |
over values of type @{typ nat}. The transfer principle does |
|
431 |
not handle quantification over non-star types in general, |
|
432 |
but we can work around this by fixing an arbitrary @{typ nat} |
|
433 |
value, and then applying the transfer principle. |
|
434 |
*} |
|
435 |
||
436 |
instance star :: (recpower) recpower |
|
437 |
proof |
|
438 |
show "\<And>a::'a star. a ^ 0 = 1" |
|
439 |
by transfer (rule power_0) |
|
440 |
next |
|
441 |
fix n show "\<And>a::'a star. a ^ Suc n = a * a ^ n" |
|
442 |
by transfer (rule power_Suc) |
|
443 |
qed |
|
444 |
||
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
445 |
subsection {* Number classes *} |
17296 | 446 |
|
20720 | 447 |
lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" |
448 |
by (induct_tac n, simp_all) |
|
449 |
||
450 |
lemma Standard_of_nat [simp]: "of_nat n \<in> Standard" |
|
451 |
by (simp add: star_of_nat_def) |
|
17296 | 452 |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
453 |
lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
454 |
by transfer (rule refl) |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
455 |
|
20720 | 456 |
lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)" |
457 |
by (rule_tac z=z in int_diff_cases, simp) |
|
458 |
||
459 |
lemma Standard_of_int [simp]: "of_int z \<in> Standard" |
|
460 |
by (simp add: star_of_int_def) |
|
17332
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
461 |
|
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
462 |
lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" |
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents:
17296
diff
changeset
|
463 |
by transfer (rule refl) |
17296 | 464 |
|
23282
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
huffman
parents:
22993
diff
changeset
|
465 |
instance star :: (semiring_char_0) semiring_char_0 |
24195 | 466 |
by intro_classes (simp only: star_of_nat_def star_of_eq of_nat_eq_iff) |
23282
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
huffman
parents:
22993
diff
changeset
|
467 |
|
dfc459989d24
add axclass semiring_char_0 for types where of_nat is injective
huffman
parents:
22993
diff
changeset
|
468 |
instance star :: (ring_char_0) ring_char_0 .. |
22911
2f5e8d70a179
new axclass ring_char_0 for rings with characteristic 0, used for of_int_eq_iff and related lemmas
huffman
parents:
22518
diff
changeset
|
469 |
|
17296 | 470 |
instance star :: (number_ring) number_ring |
471 |
by (intro_classes, simp only: star_number_def star_of_int_def number_of_eq) |
|
472 |
||
17429
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
473 |
subsection {* Finite class *} |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
474 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
475 |
lemma starset_finite: "finite A \<Longrightarrow> *s* A = star_of ` A" |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
476 |
by (erule finite_induct, simp_all) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
477 |
|
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
478 |
instance star :: (finite) finite |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
479 |
apply (intro_classes) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
480 |
apply (subst starset_UNIV [symmetric]) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
481 |
apply (subst starset_finite [OF finite]) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
482 |
apply (rule finite_imageI [OF finite]) |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
483 |
done |
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
huffman
parents:
17332
diff
changeset
|
484 |
|
17296 | 485 |
end |