| author | blanchet |
| Wed, 28 Apr 2010 18:11:11 +0200 | |
| changeset 36550 | f8da913b6c3a |
| parent 35416 | d8d7d1b785af |
| child 43786 | fea3ed6951e3 |
| permissions | -rw-r--r-- |
| 17644 | 1 |
(* AUTOMATICALLY GENERATED, DO NOT EDIT! *) |
2 |
||
| 19093 | 3 |
theory HOLLight = "../HOLLightCompat" + "../HOL4Syntax": |
| 17644 | 4 |
|
5 |
;setup_theory hollight |
|
6 |
||
7 |
consts |
|
8 |
"_FALSITY_" :: "bool" ("'_FALSITY'_")
|
|
9 |
||
10 |
defs |
|
11 |
"_FALSITY__def": "_FALSITY_ == False" |
|
12 |
||
13 |
lemma DEF__FALSITY_: "_FALSITY_ = False" |
|
14 |
by (import hollight DEF__FALSITY_) |
|
15 |
||
16 |
lemma CONJ_ACI: "((p::bool) & (q::bool)) = (q & p) & |
|
17 |
((p & q) & (r::bool)) = (p & q & r) & |
|
18 |
(p & q & r) = (q & p & r) & (p & p) = p & (p & p & q) = (p & q)" |
|
19 |
by (import hollight CONJ_ACI) |
|
20 |
||
21 |
lemma DISJ_ACI: "((p::bool) | (q::bool)) = (q | p) & |
|
22 |
((p | q) | (r::bool)) = (p | q | r) & |
|
23 |
(p | q | r) = (q | p | r) & (p | p) = p & (p | p | q) = (p | q)" |
|
24 |
by (import hollight DISJ_ACI) |
|
25 |
||
26 |
lemma EQ_CLAUSES: "ALL t::bool. |
|
27 |
(True = t) = t & |
|
28 |
(t = True) = t & (False = t) = (~ t) & (t = False) = (~ t)" |
|
29 |
by (import hollight EQ_CLAUSES) |
|
30 |
||
31 |
lemma NOT_CLAUSES_WEAK: "(~ True) = False & (~ False) = True" |
|
32 |
by (import hollight NOT_CLAUSES_WEAK) |
|
33 |
||
34 |
lemma AND_CLAUSES: "ALL t::bool. |
|
35 |
(True & t) = t & |
|
36 |
(t & True) = t & (False & t) = False & (t & False) = False & (t & t) = t" |
|
37 |
by (import hollight AND_CLAUSES) |
|
38 |
||
39 |
lemma OR_CLAUSES: "ALL t::bool. |
|
40 |
(True | t) = True & |
|
41 |
(t | True) = True & (False | t) = t & (t | False) = t & (t | t) = t" |
|
42 |
by (import hollight OR_CLAUSES) |
|
43 |
||
44 |
lemma IMP_CLAUSES: "ALL t::bool. |
|
45 |
(True --> t) = t & |
|
46 |
(t --> True) = True & |
|
47 |
(False --> t) = True & (t --> t) = True & (t --> False) = (~ t)" |
|
48 |
by (import hollight IMP_CLAUSES) |
|
49 |
||
50 |
lemma IMP_EQ_CLAUSE: "((x::'q_864::type) = x --> (p::bool)) = p" |
|
51 |
by (import hollight IMP_EQ_CLAUSE) |
|
52 |
||
53 |
lemma SWAP_FORALL_THM: "ALL P::'A::type => 'B::type => bool. |
|
54 |
(ALL x::'A::type. All (P x)) = (ALL (y::'B::type) x::'A::type. P x y)" |
|
55 |
by (import hollight SWAP_FORALL_THM) |
|
56 |
||
57 |
lemma SWAP_EXISTS_THM: "ALL P::'A::type => 'B::type => bool. |
|
58 |
(EX x::'A::type. Ex (P x)) = (EX (x::'B::type) xa::'A::type. P xa x)" |
|
59 |
by (import hollight SWAP_EXISTS_THM) |
|
60 |
||
61 |
lemma TRIV_EXISTS_AND_THM: "ALL (P::bool) Q::bool. |
|
62 |
(EX x::'A::type. P & Q) = ((EX x::'A::type. P) & (EX x::'A::type. Q))" |
|
63 |
by (import hollight TRIV_EXISTS_AND_THM) |
|
64 |
||
65 |
lemma TRIV_AND_EXISTS_THM: "ALL (P::bool) Q::bool. |
|
66 |
((EX x::'A::type. P) & (EX x::'A::type. Q)) = (EX x::'A::type. P & Q)" |
|
67 |
by (import hollight TRIV_AND_EXISTS_THM) |
|
68 |
||
69 |
lemma TRIV_FORALL_OR_THM: "ALL (P::bool) Q::bool. |
|
70 |
(ALL x::'A::type. P | Q) = ((ALL x::'A::type. P) | (ALL x::'A::type. Q))" |
|
71 |
by (import hollight TRIV_FORALL_OR_THM) |
|
72 |
||
73 |
lemma TRIV_OR_FORALL_THM: "ALL (P::bool) Q::bool. |
|
74 |
((ALL x::'A::type. P) | (ALL x::'A::type. Q)) = (ALL x::'A::type. P | Q)" |
|
75 |
by (import hollight TRIV_OR_FORALL_THM) |
|
76 |
||
77 |
lemma TRIV_FORALL_IMP_THM: "ALL (P::bool) Q::bool. |
|
78 |
(ALL x::'A::type. P --> Q) = |
|
79 |
((EX x::'A::type. P) --> (ALL x::'A::type. Q))" |
|
80 |
by (import hollight TRIV_FORALL_IMP_THM) |
|
81 |
||
82 |
lemma TRIV_EXISTS_IMP_THM: "ALL (P::bool) Q::bool. |
|
83 |
(EX x::'A::type. P --> Q) = |
|
84 |
((ALL x::'A::type. P) --> (EX x::'A::type. Q))" |
|
85 |
by (import hollight TRIV_EXISTS_IMP_THM) |
|
86 |
||
87 |
lemma EXISTS_UNIQUE_ALT: "ALL P::'A::type => bool. |
|
88 |
Ex1 P = (EX x::'A::type. ALL y::'A::type. P y = (x = y))" |
|
89 |
by (import hollight EXISTS_UNIQUE_ALT) |
|
90 |
||
91 |
lemma SELECT_UNIQUE: "ALL (P::'A::type => bool) x::'A::type. |
|
92 |
(ALL y::'A::type. P y = (y = x)) --> Eps P = x" |
|
93 |
by (import hollight SELECT_UNIQUE) |
|
94 |
||
95 |
lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t" |
|
96 |
by (import hollight EXCLUDED_MIDDLE) |
|
97 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
98 |
definition COND :: "bool => 'A => 'A => 'A" where |
| 17644 | 99 |
"COND == |
100 |
%(t::bool) (t1::'A::type) t2::'A::type. |
|
101 |
SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2)" |
|
102 |
||
103 |
lemma DEF_COND: "COND = |
|
104 |
(%(t::bool) (t1::'A::type) t2::'A::type. |
|
105 |
SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2))" |
|
106 |
by (import hollight DEF_COND) |
|
107 |
||
108 |
lemma COND_CLAUSES: "ALL (x::'A::type) xa::'A::type. COND True x xa = x & COND False x xa = xa" |
|
109 |
by (import hollight COND_CLAUSES) |
|
110 |
||
111 |
lemma COND_EXPAND: "ALL (b::bool) (t1::bool) t2::bool. COND b t1 t2 = ((~ b | t1) & (b | t2))" |
|
112 |
by (import hollight COND_EXPAND) |
|
113 |
||
114 |
lemma COND_ID: "ALL (b::bool) t::'A::type. COND b t t = t" |
|
115 |
by (import hollight COND_ID) |
|
116 |
||
117 |
lemma COND_RAND: "ALL (b::bool) (f::'A::type => 'B::type) (x::'A::type) y::'A::type. |
|
118 |
f (COND b x y) = COND b (f x) (f y)" |
|
119 |
by (import hollight COND_RAND) |
|
120 |
||
121 |
lemma COND_RATOR: "ALL (b::bool) (f::'A::type => 'B::type) (g::'A::type => 'B::type) |
|
122 |
x::'A::type. COND b f g x = COND b (f x) (g x)" |
|
123 |
by (import hollight COND_RATOR) |
|
124 |
||
125 |
lemma COND_ABS: "ALL (b::bool) (f::'A::type => 'B::type) g::'A::type => 'B::type. |
|
126 |
(%x::'A::type. COND b (f x) (g x)) = COND b f g" |
|
127 |
by (import hollight COND_ABS) |
|
128 |
||
129 |
lemma MONO_COND: "((A::bool) --> (B::bool)) & ((C::bool) --> (D::bool)) --> |
|
130 |
COND (b::bool) A C --> COND b B D" |
|
131 |
by (import hollight MONO_COND) |
|
132 |
||
133 |
lemma COND_ELIM_THM: "(P::'A::type => bool) (COND (c::bool) (x::'A::type) (y::'A::type)) = |
|
134 |
((c --> P x) & (~ c --> P y))" |
|
135 |
by (import hollight COND_ELIM_THM) |
|
136 |
||
137 |
lemma SKOLEM_THM: "ALL P::'A::type => 'B::type => bool. |
|
138 |
(ALL x::'A::type. Ex (P x)) = |
|
139 |
(EX x::'A::type => 'B::type. ALL xa::'A::type. P xa (x xa))" |
|
140 |
by (import hollight SKOLEM_THM) |
|
141 |
||
142 |
lemma UNIQUE_SKOLEM_ALT: "ALL P::'A::type => 'B::type => bool. |
|
143 |
(ALL x::'A::type. Ex1 (P x)) = |
|
144 |
(EX f::'A::type => 'B::type. |
|
145 |
ALL (x::'A::type) y::'B::type. P x y = (f x = y))" |
|
146 |
by (import hollight UNIQUE_SKOLEM_ALT) |
|
147 |
||
148 |
lemma COND_EQ_CLAUSE: "COND ((x::'q_3000::type) = x) (y::'q_2993::type) (z::'q_2993::type) = y" |
|
149 |
by (import hollight COND_EQ_CLAUSE) |
|
150 |
||
151 |
lemma o_ASSOC: "ALL (f::'C::type => 'D::type) (g::'B::type => 'C::type) |
|
152 |
h::'A::type => 'B::type. f o (g o h) = f o g o h" |
|
153 |
by (import hollight o_ASSOC) |
|
154 |
||
155 |
lemma I_O_ID: "ALL f::'A::type => 'B::type. id o f = f & f o id = f" |
|
156 |
by (import hollight I_O_ID) |
|
157 |
||
158 |
lemma EXISTS_ONE_REP: "EX x::bool. x" |
|
159 |
by (import hollight EXISTS_ONE_REP) |
|
160 |
||
161 |
lemma one_axiom: "ALL f::'A::type => unit. All (op = f)" |
|
162 |
by (import hollight one_axiom) |
|
163 |
||
164 |
lemma one_RECURSION: "ALL e::'A::type. EX x::unit => 'A::type. x () = e" |
|
165 |
by (import hollight one_RECURSION) |
|
166 |
||
167 |
lemma one_Axiom: "ALL e::'A::type. EX! fn::unit => 'A::type. fn () = e" |
|
168 |
by (import hollight one_Axiom) |
|
169 |
||
170 |
lemma th_cond: "(P::'A::type => bool => bool) (COND (b::bool) (x::'A::type) (y::'A::type)) |
|
171 |
b = |
|
172 |
(b & P x True | ~ b & P y False)" |
|
173 |
by (import hollight th_cond) |
|
174 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
175 |
definition LET_END :: "'A => 'A" where |
| 17644 | 176 |
"LET_END == %t::'A::type. t" |
177 |
||
178 |
lemma DEF_LET_END: "LET_END = (%t::'A::type. t)" |
|
179 |
by (import hollight DEF_LET_END) |
|
180 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
181 |
definition GABS :: "('A => bool) => 'A" where
|
| 17644 | 182 |
"(op ==::(('A::type => bool) => 'A::type)
|
183 |
=> (('A::type => bool) => 'A::type) => prop)
|
|
184 |
(GABS::('A::type => bool) => 'A::type)
|
|
185 |
(Eps::('A::type => bool) => 'A::type)"
|
|
186 |
||
187 |
lemma DEF_GABS: "(op =::(('A::type => bool) => 'A::type)
|
|
188 |
=> (('A::type => bool) => 'A::type) => bool)
|
|
189 |
(GABS::('A::type => bool) => 'A::type)
|
|
190 |
(Eps::('A::type => bool) => 'A::type)"
|
|
191 |
by (import hollight DEF_GABS) |
|
192 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
193 |
definition GEQ :: "'A => 'A => bool" where |
| 17644 | 194 |
"(op ==::('A::type => 'A::type => bool)
|
195 |
=> ('A::type => 'A::type => bool) => prop)
|
|
196 |
(GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)" |
|
197 |
||
198 |
lemma DEF_GEQ: "(op =::('A::type => 'A::type => bool)
|
|
199 |
=> ('A::type => 'A::type => bool) => bool)
|
|
200 |
(GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)" |
|
201 |
by (import hollight DEF_GEQ) |
|
202 |
||
203 |
lemma PAIR_EXISTS_THM: "EX (x::'A::type => 'B::type => bool) (a::'A::type) b::'B::type. |
|
204 |
x = Pair_Rep a b" |
|
205 |
by (import hollight PAIR_EXISTS_THM) |
|
206 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
207 |
definition CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" where
|
| 17644 | 208 |
"CURRY == |
209 |
%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type. |
|
210 |
u (ua, ub)" |
|
211 |
||
212 |
lemma DEF_CURRY: "CURRY = |
|
213 |
(%(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type. |
|
214 |
u (ua, ub))" |
|
215 |
by (import hollight DEF_CURRY) |
|
216 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
217 |
definition UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where
|
| 17644 | 218 |
"UNCURRY == |
219 |
%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type. |
|
220 |
u (fst ua) (snd ua)" |
|
221 |
||
222 |
lemma DEF_UNCURRY: "UNCURRY = |
|
223 |
(%(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type. |
|
224 |
u (fst ua) (snd ua))" |
|
225 |
by (import hollight DEF_UNCURRY) |
|
226 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
227 |
definition PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where
|
| 17644 | 228 |
"PASSOC == |
229 |
%(u::('A::type * 'B::type) * 'C::type => 'D::type)
|
|
230 |
ua::'A::type * 'B::type * 'C::type. |
|
231 |
u ((fst ua, fst (snd ua)), snd (snd ua))" |
|
232 |
||
233 |
lemma DEF_PASSOC: "PASSOC = |
|
234 |
(%(u::('A::type * 'B::type) * 'C::type => 'D::type)
|
|
235 |
ua::'A::type * 'B::type * 'C::type. |
|
236 |
u ((fst ua, fst (snd ua)), snd (snd ua)))" |
|
237 |
by (import hollight DEF_PASSOC) |
|
238 |
||
239 |
lemma num_Axiom: "ALL (e::'A::type) f::'A::type => nat => 'A::type. |
|
| 17652 | 240 |
EX! fn::nat => 'A::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f (fn n) n)" |
| 17644 | 241 |
by (import hollight num_Axiom) |
242 |
||
| 17652 | 243 |
lemma ADD_CLAUSES: "(ALL x::nat. 0 + x = x) & |
244 |
(ALL x::nat. x + 0 = x) & |
|
| 17644 | 245 |
(ALL (x::nat) xa::nat. Suc x + xa = Suc (x + xa)) & |
246 |
(ALL (x::nat) xa::nat. x + Suc xa = Suc (x + xa))" |
|
247 |
by (import hollight ADD_CLAUSES) |
|
248 |
||
249 |
lemma ADD_AC: "(m::nat) + (n::nat) = n + m & |
|
250 |
m + n + (p::nat) = m + (n + p) & m + (n + p) = n + (m + p)" |
|
251 |
by (import hollight ADD_AC) |
|
252 |
||
| 17652 | 253 |
lemma EQ_ADD_LCANCEL_0: "ALL (m::nat) n::nat. (m + n = m) = (n = 0)" |
| 17644 | 254 |
by (import hollight EQ_ADD_LCANCEL_0) |
255 |
||
| 17652 | 256 |
lemma EQ_ADD_RCANCEL_0: "ALL (x::nat) xa::nat. (x + xa = xa) = (x = 0)" |
| 17644 | 257 |
by (import hollight EQ_ADD_RCANCEL_0) |
258 |
||
| 17652 | 259 |
lemma ONE: "NUMERAL_BIT1 0 = Suc 0" |
| 17644 | 260 |
by (import hollight ONE) |
261 |
||
| 17652 | 262 |
lemma TWO: "NUMERAL_BIT0 (NUMERAL_BIT1 0) = Suc (NUMERAL_BIT1 0)" |
| 17644 | 263 |
by (import hollight TWO) |
264 |
||
| 17652 | 265 |
lemma ADD1: "ALL x::nat. Suc x = x + NUMERAL_BIT1 0" |
| 17644 | 266 |
by (import hollight ADD1) |
267 |
||
| 17652 | 268 |
lemma MULT_CLAUSES: "(ALL x::nat. 0 * x = 0) & |
269 |
(ALL x::nat. x * 0 = 0) & |
|
270 |
(ALL x::nat. NUMERAL_BIT1 0 * x = x) & |
|
271 |
(ALL x::nat. x * NUMERAL_BIT1 0 = x) & |
|
| 17644 | 272 |
(ALL (x::nat) xa::nat. Suc x * xa = x * xa + xa) & |
273 |
(ALL (x::nat) xa::nat. x * Suc xa = x + x * xa)" |
|
274 |
by (import hollight MULT_CLAUSES) |
|
275 |
||
276 |
lemma MULT_AC: "(m::nat) * (n::nat) = n * m & |
|
277 |
m * n * (p::nat) = m * (n * p) & m * (n * p) = n * (m * p)" |
|
278 |
by (import hollight MULT_AC) |
|
279 |
||
| 17652 | 280 |
lemma MULT_2: "ALL n::nat. NUMERAL_BIT0 (NUMERAL_BIT1 0) * n = n + n" |
| 17644 | 281 |
by (import hollight MULT_2) |
282 |
||
283 |
lemma MULT_EQ_1: "ALL (m::nat) n::nat. |
|
| 17652 | 284 |
(m * n = NUMERAL_BIT1 0) = (m = NUMERAL_BIT1 0 & n = NUMERAL_BIT1 0)" |
| 17644 | 285 |
by (import hollight MULT_EQ_1) |
286 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
287 |
definition EXP :: "nat => nat => nat" where |
| 17644 | 288 |
"EXP == |
289 |
SOME EXP::nat => nat => nat. |
|
| 17652 | 290 |
(ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) & |
| 17644 | 291 |
(ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n)" |
292 |
||
293 |
lemma DEF_EXP: "EXP = |
|
294 |
(SOME EXP::nat => nat => nat. |
|
| 17652 | 295 |
(ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) & |
| 17644 | 296 |
(ALL (m::nat) n::nat. EXP m (Suc n) = m * EXP m n))" |
297 |
by (import hollight DEF_EXP) |
|
298 |
||
| 17652 | 299 |
lemma EXP_EQ_0: "ALL (m::nat) n::nat. (EXP m n = 0) = (m = 0 & n ~= 0)" |
| 17644 | 300 |
by (import hollight EXP_EQ_0) |
301 |
||
302 |
lemma EXP_ADD: "ALL (m::nat) (n::nat) p::nat. EXP m (n + p) = EXP m n * EXP m p" |
|
303 |
by (import hollight EXP_ADD) |
|
304 |
||
| 17652 | 305 |
lemma EXP_ONE: "ALL n::nat. EXP (NUMERAL_BIT1 0) n = NUMERAL_BIT1 0" |
| 17644 | 306 |
by (import hollight EXP_ONE) |
307 |
||
| 17652 | 308 |
lemma EXP_1: "ALL x::nat. EXP x (NUMERAL_BIT1 0) = x" |
| 17644 | 309 |
by (import hollight EXP_1) |
310 |
||
| 17652 | 311 |
lemma EXP_2: "ALL x::nat. EXP x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = x * x" |
| 17644 | 312 |
by (import hollight EXP_2) |
313 |
||
314 |
lemma MULT_EXP: "ALL (p::nat) (m::nat) n::nat. EXP (m * n) p = EXP m p * EXP n p" |
|
315 |
by (import hollight MULT_EXP) |
|
316 |
||
317 |
lemma EXP_MULT: "ALL (m::nat) (n::nat) p::nat. EXP m (n * p) = EXP (EXP m n) p" |
|
318 |
by (import hollight EXP_MULT) |
|
319 |
||
320 |
consts |
|
321 |
"<=" :: "nat => nat => bool" ("<=")
|
|
322 |
||
323 |
defs |
|
324 |
"<=_def": "<= == |
|
325 |
SOME u::nat => nat => bool. |
|
| 17652 | 326 |
(ALL m::nat. u m 0 = (m = 0)) & |
| 17644 | 327 |
(ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n))" |
328 |
||
329 |
lemma DEF__lessthan__equal_: "<= = |
|
330 |
(SOME u::nat => nat => bool. |
|
| 17652 | 331 |
(ALL m::nat. u m 0 = (m = 0)) & |
| 17644 | 332 |
(ALL (m::nat) n::nat. u m (Suc n) = (m = Suc n | u m n)))" |
333 |
by (import hollight DEF__lessthan__equal_) |
|
334 |
||
335 |
consts |
|
336 |
"<" :: "nat => nat => bool" ("<")
|
|
337 |
||
338 |
defs |
|
339 |
"<_def": "< == |
|
340 |
SOME u::nat => nat => bool. |
|
| 17652 | 341 |
(ALL m::nat. u m 0 = False) & |
| 17644 | 342 |
(ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n))" |
343 |
||
344 |
lemma DEF__lessthan_: "< = |
|
345 |
(SOME u::nat => nat => bool. |
|
| 17652 | 346 |
(ALL m::nat. u m 0 = False) & |
| 17644 | 347 |
(ALL (m::nat) n::nat. u m (Suc n) = (m = n | u m n)))" |
348 |
by (import hollight DEF__lessthan_) |
|
349 |
||
350 |
consts |
|
351 |
">=" :: "nat => nat => bool" (">=")
|
|
352 |
||
353 |
defs |
|
354 |
">=_def": ">= == %(u::nat) ua::nat. <= ua u" |
|
355 |
||
356 |
lemma DEF__greaterthan__equal_: ">= = (%(u::nat) ua::nat. <= ua u)" |
|
357 |
by (import hollight DEF__greaterthan__equal_) |
|
358 |
||
359 |
consts |
|
360 |
">" :: "nat => nat => bool" (">")
|
|
361 |
||
362 |
defs |
|
363 |
">_def": "> == %(u::nat) ua::nat. < ua u" |
|
364 |
||
365 |
lemma DEF__greaterthan_: "> = (%(u::nat) ua::nat. < ua u)" |
|
366 |
by (import hollight DEF__greaterthan_) |
|
367 |
||
368 |
lemma LE_SUC_LT: "ALL (m::nat) n::nat. <= (Suc m) n = < m n" |
|
369 |
by (import hollight LE_SUC_LT) |
|
370 |
||
371 |
lemma LT_SUC_LE: "ALL (m::nat) n::nat. < m (Suc n) = <= m n" |
|
372 |
by (import hollight LT_SUC_LE) |
|
373 |
||
374 |
lemma LE_SUC: "ALL (x::nat) xa::nat. <= (Suc x) (Suc xa) = <= x xa" |
|
375 |
by (import hollight LE_SUC) |
|
376 |
||
377 |
lemma LT_SUC: "ALL (x::nat) xa::nat. < (Suc x) (Suc xa) = < x xa" |
|
378 |
by (import hollight LT_SUC) |
|
379 |
||
| 17652 | 380 |
lemma LE_0: "All (<= 0)" |
| 17644 | 381 |
by (import hollight LE_0) |
382 |
||
| 17652 | 383 |
lemma LT_0: "ALL x::nat. < 0 (Suc x)" |
| 17644 | 384 |
by (import hollight LT_0) |
385 |
||
386 |
lemma LE_REFL: "ALL n::nat. <= n n" |
|
387 |
by (import hollight LE_REFL) |
|
388 |
||
389 |
lemma LT_REFL: "ALL n::nat. ~ < n n" |
|
390 |
by (import hollight LT_REFL) |
|
391 |
||
392 |
lemma LE_ANTISYM: "ALL (m::nat) n::nat. (<= m n & <= n m) = (m = n)" |
|
393 |
by (import hollight LE_ANTISYM) |
|
394 |
||
395 |
lemma LT_ANTISYM: "ALL (m::nat) n::nat. ~ (< m n & < n m)" |
|
396 |
by (import hollight LT_ANTISYM) |
|
397 |
||
398 |
lemma LET_ANTISYM: "ALL (m::nat) n::nat. ~ (<= m n & < n m)" |
|
399 |
by (import hollight LET_ANTISYM) |
|
400 |
||
401 |
lemma LTE_ANTISYM: "ALL (x::nat) xa::nat. ~ (< x xa & <= xa x)" |
|
402 |
by (import hollight LTE_ANTISYM) |
|
403 |
||
404 |
lemma LE_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & <= n p --> <= m p" |
|
405 |
by (import hollight LE_TRANS) |
|
406 |
||
407 |
lemma LT_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & < n p --> < m p" |
|
408 |
by (import hollight LT_TRANS) |
|
409 |
||
410 |
lemma LET_TRANS: "ALL (m::nat) (n::nat) p::nat. <= m n & < n p --> < m p" |
|
411 |
by (import hollight LET_TRANS) |
|
412 |
||
413 |
lemma LTE_TRANS: "ALL (m::nat) (n::nat) p::nat. < m n & <= n p --> < m p" |
|
414 |
by (import hollight LTE_TRANS) |
|
415 |
||
416 |
lemma LE_CASES: "ALL (m::nat) n::nat. <= m n | <= n m" |
|
417 |
by (import hollight LE_CASES) |
|
418 |
||
419 |
lemma LT_CASES: "ALL (m::nat) n::nat. < m n | < n m | m = n" |
|
420 |
by (import hollight LT_CASES) |
|
421 |
||
422 |
lemma LET_CASES: "ALL (m::nat) n::nat. <= m n | < n m" |
|
423 |
by (import hollight LET_CASES) |
|
424 |
||
425 |
lemma LTE_CASES: "ALL (x::nat) xa::nat. < x xa | <= xa x" |
|
426 |
by (import hollight LTE_CASES) |
|
427 |
||
| 17652 | 428 |
lemma LT_NZ: "ALL n::nat. < 0 n = (n ~= 0)" |
| 17644 | 429 |
by (import hollight LT_NZ) |
430 |
||
431 |
lemma LE_LT: "ALL (m::nat) n::nat. <= m n = (< m n | m = n)" |
|
432 |
by (import hollight LE_LT) |
|
433 |
||
434 |
lemma LT_LE: "ALL (x::nat) xa::nat. < x xa = (<= x xa & x ~= xa)" |
|
435 |
by (import hollight LT_LE) |
|
436 |
||
437 |
lemma NOT_LE: "ALL (m::nat) n::nat. (~ <= m n) = < n m" |
|
438 |
by (import hollight NOT_LE) |
|
439 |
||
440 |
lemma NOT_LT: "ALL (m::nat) n::nat. (~ < m n) = <= n m" |
|
441 |
by (import hollight NOT_LT) |
|
442 |
||
443 |
lemma LT_IMP_LE: "ALL (x::nat) xa::nat. < x xa --> <= x xa" |
|
444 |
by (import hollight LT_IMP_LE) |
|
445 |
||
446 |
lemma EQ_IMP_LE: "ALL (m::nat) n::nat. m = n --> <= m n" |
|
447 |
by (import hollight EQ_IMP_LE) |
|
448 |
||
449 |
lemma LE_EXISTS: "ALL (m::nat) n::nat. <= m n = (EX d::nat. n = m + d)" |
|
450 |
by (import hollight LE_EXISTS) |
|
451 |
||
452 |
lemma LT_EXISTS: "ALL (m::nat) n::nat. < m n = (EX d::nat. n = m + Suc d)" |
|
453 |
by (import hollight LT_EXISTS) |
|
454 |
||
455 |
lemma LE_ADD: "ALL (m::nat) n::nat. <= m (m + n)" |
|
456 |
by (import hollight LE_ADD) |
|
457 |
||
458 |
lemma LE_ADDR: "ALL (x::nat) xa::nat. <= xa (x + xa)" |
|
459 |
by (import hollight LE_ADDR) |
|
460 |
||
| 17652 | 461 |
lemma LT_ADD: "ALL (m::nat) n::nat. < m (m + n) = < 0 n" |
| 17644 | 462 |
by (import hollight LT_ADD) |
463 |
||
| 17652 | 464 |
lemma LT_ADDR: "ALL (x::nat) xa::nat. < xa (x + xa) = < 0 x" |
| 17644 | 465 |
by (import hollight LT_ADDR) |
466 |
||
467 |
lemma LE_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xa) (x + xb) = <= xa xb" |
|
468 |
by (import hollight LE_ADD_LCANCEL) |
|
469 |
||
470 |
lemma LE_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x + xb) (xa + xb) = <= x xa" |
|
471 |
by (import hollight LE_ADD_RCANCEL) |
|
472 |
||
473 |
lemma LT_ADD_LCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xa) (x + xb) = < xa xb" |
|
474 |
by (import hollight LT_ADD_LCANCEL) |
|
475 |
||
476 |
lemma LT_ADD_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x + xb) (xa + xb) = < x xa" |
|
477 |
by (import hollight LT_ADD_RCANCEL) |
|
478 |
||
479 |
lemma LE_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
480 |
<= m p & <= n q --> <= (m + n) (p + q)" |
|
481 |
by (import hollight LE_ADD2) |
|
482 |
||
483 |
lemma LET_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. <= m p & < n q --> < (m + n) (p + q)" |
|
484 |
by (import hollight LET_ADD2) |
|
485 |
||
486 |
lemma LTE_ADD2: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. |
|
487 |
< x xb & <= xa xc --> < (x + xa) (xb + xc)" |
|
488 |
by (import hollight LTE_ADD2) |
|
489 |
||
490 |
lemma LT_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m p & < n q --> < (m + n) (p + q)" |
|
491 |
by (import hollight LT_ADD2) |
|
492 |
||
| 17652 | 493 |
lemma LT_MULT: "ALL (m::nat) n::nat. < 0 (m * n) = (< 0 m & < 0 n)" |
| 17644 | 494 |
by (import hollight LT_MULT) |
495 |
||
496 |
lemma LE_MULT2: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
497 |
<= m n & <= p q --> <= (m * p) (n * q)" |
|
498 |
by (import hollight LE_MULT2) |
|
499 |
||
| 17652 | 500 |
lemma LT_LMULT: "ALL (m::nat) (n::nat) p::nat. m ~= 0 & < n p --> < (m * n) (m * p)" |
| 17644 | 501 |
by (import hollight LT_LMULT) |
502 |
||
| 17652 | 503 |
lemma LE_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. <= (m * n) (m * p) = (m = 0 | <= n p)" |
| 17644 | 504 |
by (import hollight LE_MULT_LCANCEL) |
505 |
||
| 17652 | 506 |
lemma LE_MULT_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. <= (x * xb) (xa * xb) = (<= x xa | xb = 0)" |
| 17644 | 507 |
by (import hollight LE_MULT_RCANCEL) |
508 |
||
| 17652 | 509 |
lemma LT_MULT_LCANCEL: "ALL (m::nat) (n::nat) p::nat. < (m * n) (m * p) = (m ~= 0 & < n p)" |
| 17644 | 510 |
by (import hollight LT_MULT_LCANCEL) |
511 |
||
| 17652 | 512 |
lemma LT_MULT_RCANCEL: "ALL (x::nat) (xa::nat) xb::nat. < (x * xb) (xa * xb) = (< x xa & xb ~= 0)" |
| 17644 | 513 |
by (import hollight LT_MULT_RCANCEL) |
514 |
||
515 |
lemma LT_MULT2: "ALL (m::nat) (n::nat) (p::nat) q::nat. < m n & < p q --> < (m * p) (n * q)" |
|
516 |
by (import hollight LT_MULT2) |
|
517 |
||
518 |
lemma LE_SQUARE_REFL: "ALL n::nat. <= n (n * n)" |
|
519 |
by (import hollight LE_SQUARE_REFL) |
|
520 |
||
521 |
lemma WLOG_LE: "(ALL (m::nat) n::nat. (P::nat => nat => bool) m n = P n m) & |
|
522 |
(ALL (m::nat) n::nat. <= m n --> P m n) --> |
|
523 |
(ALL m::nat. All (P m))" |
|
524 |
by (import hollight WLOG_LE) |
|
525 |
||
526 |
lemma WLOG_LT: "(ALL m::nat. (P::nat => nat => bool) m m) & |
|
527 |
(ALL (m::nat) n::nat. P m n = P n m) & |
|
528 |
(ALL (m::nat) n::nat. < m n --> P m n) --> |
|
529 |
(ALL m::nat. All (P m))" |
|
530 |
by (import hollight WLOG_LT) |
|
531 |
||
532 |
lemma num_WF: "ALL P::nat => bool. |
|
533 |
(ALL n::nat. (ALL m::nat. < m n --> P m) --> P n) --> All P" |
|
534 |
by (import hollight num_WF) |
|
535 |
||
536 |
lemma num_WOP: "ALL P::nat => bool. Ex P = (EX n::nat. P n & (ALL m::nat. < m n --> ~ P m))" |
|
537 |
by (import hollight num_WOP) |
|
538 |
||
539 |
lemma num_MAX: "ALL P::nat => bool. |
|
540 |
(Ex P & (EX M::nat. ALL x::nat. P x --> <= x M)) = |
|
541 |
(EX m::nat. P m & (ALL x::nat. P x --> <= x m))" |
|
542 |
by (import hollight num_MAX) |
|
543 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
544 |
definition EVEN :: "nat => bool" where |
| 17644 | 545 |
"EVEN == |
546 |
SOME EVEN::nat => bool. |
|
| 17652 | 547 |
EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))" |
| 17644 | 548 |
|
549 |
lemma DEF_EVEN: "EVEN = |
|
550 |
(SOME EVEN::nat => bool. |
|
| 17652 | 551 |
EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))" |
| 17644 | 552 |
by (import hollight DEF_EVEN) |
553 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
554 |
definition ODD :: "nat => bool" where |
| 17644 | 555 |
"ODD == |
| 17652 | 556 |
SOME ODD::nat => bool. ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))" |
| 17644 | 557 |
|
558 |
lemma DEF_ODD: "ODD = |
|
559 |
(SOME ODD::nat => bool. |
|
| 17652 | 560 |
ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n)))" |
| 17644 | 561 |
by (import hollight DEF_ODD) |
562 |
||
563 |
lemma NOT_EVEN: "ALL n::nat. (~ EVEN n) = ODD n" |
|
564 |
by (import hollight NOT_EVEN) |
|
565 |
||
566 |
lemma NOT_ODD: "ALL n::nat. (~ ODD n) = EVEN n" |
|
567 |
by (import hollight NOT_ODD) |
|
568 |
||
569 |
lemma EVEN_OR_ODD: "ALL n::nat. EVEN n | ODD n" |
|
570 |
by (import hollight EVEN_OR_ODD) |
|
571 |
||
572 |
lemma EVEN_AND_ODD: "ALL x::nat. ~ (EVEN x & ODD x)" |
|
573 |
by (import hollight EVEN_AND_ODD) |
|
574 |
||
575 |
lemma EVEN_ADD: "ALL (m::nat) n::nat. EVEN (m + n) = (EVEN m = EVEN n)" |
|
576 |
by (import hollight EVEN_ADD) |
|
577 |
||
578 |
lemma EVEN_MULT: "ALL (m::nat) n::nat. EVEN (m * n) = (EVEN m | EVEN n)" |
|
579 |
by (import hollight EVEN_MULT) |
|
580 |
||
| 17652 | 581 |
lemma EVEN_EXP: "ALL (m::nat) n::nat. EVEN (EXP m n) = (EVEN m & n ~= 0)" |
| 17644 | 582 |
by (import hollight EVEN_EXP) |
583 |
||
584 |
lemma ODD_ADD: "ALL (m::nat) n::nat. ODD (m + n) = (ODD m ~= ODD n)" |
|
585 |
by (import hollight ODD_ADD) |
|
586 |
||
587 |
lemma ODD_MULT: "ALL (m::nat) n::nat. ODD (m * n) = (ODD m & ODD n)" |
|
588 |
by (import hollight ODD_MULT) |
|
589 |
||
| 17652 | 590 |
lemma ODD_EXP: "ALL (m::nat) n::nat. ODD (EXP m n) = (ODD m | n = 0)" |
| 17644 | 591 |
by (import hollight ODD_EXP) |
592 |
||
| 17652 | 593 |
lemma EVEN_DOUBLE: "ALL n::nat. EVEN (NUMERAL_BIT0 (NUMERAL_BIT1 0) * n)" |
| 17644 | 594 |
by (import hollight EVEN_DOUBLE) |
595 |
||
| 17652 | 596 |
lemma ODD_DOUBLE: "ALL x::nat. ODD (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * x))" |
| 17644 | 597 |
by (import hollight ODD_DOUBLE) |
598 |
||
599 |
lemma EVEN_EXISTS_LEMMA: "ALL n::nat. |
|
| 17652 | 600 |
(EVEN n --> (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)) & |
601 |
(~ EVEN n --> (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)))" |
|
| 17644 | 602 |
by (import hollight EVEN_EXISTS_LEMMA) |
603 |
||
| 17652 | 604 |
lemma EVEN_EXISTS: "ALL n::nat. EVEN n = (EX m::nat. n = NUMERAL_BIT0 (NUMERAL_BIT1 0) * m)" |
| 17644 | 605 |
by (import hollight EVEN_EXISTS) |
606 |
||
| 17652 | 607 |
lemma ODD_EXISTS: "ALL n::nat. ODD n = (EX m::nat. n = Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * m))" |
| 17644 | 608 |
by (import hollight ODD_EXISTS) |
609 |
||
610 |
lemma EVEN_ODD_DECOMPOSITION: "ALL n::nat. |
|
611 |
(EX (k::nat) m::nat. |
|
| 17652 | 612 |
ODD m & n = EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) k * m) = |
613 |
(n ~= 0)" |
|
| 17644 | 614 |
by (import hollight EVEN_ODD_DECOMPOSITION) |
615 |
||
| 17652 | 616 |
lemma SUB_0: "ALL x::nat. 0 - x = 0 & x - 0 = x" |
| 17644 | 617 |
by (import hollight SUB_0) |
618 |
||
619 |
lemma SUB_PRESUC: "ALL (m::nat) n::nat. Pred (Suc m - n) = m - n" |
|
620 |
by (import hollight SUB_PRESUC) |
|
621 |
||
| 17652 | 622 |
lemma SUB_EQ_0: "ALL (m::nat) n::nat. (m - n = 0) = <= m n" |
| 17644 | 623 |
by (import hollight SUB_EQ_0) |
624 |
||
| 17652 | 625 |
lemma ADD_SUBR: "ALL (x::nat) xa::nat. xa - (x + xa) = 0" |
| 17644 | 626 |
by (import hollight ADD_SUBR) |
627 |
||
628 |
lemma SUB_ADD: "ALL (x::nat) xa::nat. <= xa x --> x - xa + xa = x" |
|
629 |
by (import hollight SUB_ADD) |
|
630 |
||
| 17652 | 631 |
lemma SUC_SUB1: "ALL x::nat. Suc x - NUMERAL_BIT1 0 = x" |
| 17644 | 632 |
by (import hollight SUC_SUB1) |
633 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
634 |
definition FACT :: "nat => nat" where |
| 17644 | 635 |
"FACT == |
636 |
SOME FACT::nat => nat. |
|
| 17652 | 637 |
FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)" |
| 17644 | 638 |
|
639 |
lemma DEF_FACT: "FACT = |
|
640 |
(SOME FACT::nat => nat. |
|
| 17652 | 641 |
FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n))" |
| 17644 | 642 |
by (import hollight DEF_FACT) |
643 |
||
| 17652 | 644 |
lemma FACT_LT: "ALL n::nat. < 0 (FACT n)" |
| 17644 | 645 |
by (import hollight FACT_LT) |
646 |
||
| 17652 | 647 |
lemma FACT_LE: "ALL x::nat. <= (NUMERAL_BIT1 0) (FACT x)" |
| 17644 | 648 |
by (import hollight FACT_LE) |
649 |
||
650 |
lemma FACT_MONO: "ALL (m::nat) n::nat. <= m n --> <= (FACT m) (FACT n)" |
|
651 |
by (import hollight FACT_MONO) |
|
652 |
||
| 17652 | 653 |
lemma DIVMOD_EXIST: "ALL (m::nat) n::nat. n ~= 0 --> (EX (q::nat) r::nat. m = q * n + r & < r n)" |
| 17644 | 654 |
by (import hollight DIVMOD_EXIST) |
655 |
||
656 |
lemma DIVMOD_EXIST_0: "ALL (m::nat) n::nat. |
|
657 |
EX (x::nat) xa::nat. |
|
| 17652 | 658 |
COND (n = 0) (x = 0 & xa = 0) (m = x * n + xa & < xa n)" |
| 17644 | 659 |
by (import hollight DIVMOD_EXIST_0) |
660 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
661 |
definition DIV :: "nat => nat => nat" where |
| 17644 | 662 |
"DIV == |
663 |
SOME q::nat => nat => nat. |
|
664 |
EX r::nat => nat => nat. |
|
665 |
ALL (m::nat) n::nat. |
|
| 17652 | 666 |
COND (n = 0) (q m n = 0 & r m n = 0) |
| 17644 | 667 |
(m = q m n * n + r m n & < (r m n) n)" |
668 |
||
669 |
lemma DEF_DIV: "DIV = |
|
670 |
(SOME q::nat => nat => nat. |
|
671 |
EX r::nat => nat => nat. |
|
672 |
ALL (m::nat) n::nat. |
|
| 17652 | 673 |
COND (n = 0) (q m n = 0 & r m n = 0) |
| 17644 | 674 |
(m = q m n * n + r m n & < (r m n) n))" |
675 |
by (import hollight DEF_DIV) |
|
676 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
677 |
definition MOD :: "nat => nat => nat" where |
| 17644 | 678 |
"MOD == |
679 |
SOME r::nat => nat => nat. |
|
680 |
ALL (m::nat) n::nat. |
|
| 17652 | 681 |
COND (n = 0) (DIV m n = 0 & r m n = 0) |
| 17644 | 682 |
(m = DIV m n * n + r m n & < (r m n) n)" |
683 |
||
684 |
lemma DEF_MOD: "MOD = |
|
685 |
(SOME r::nat => nat => nat. |
|
686 |
ALL (m::nat) n::nat. |
|
| 17652 | 687 |
COND (n = 0) (DIV m n = 0 & r m n = 0) |
| 17644 | 688 |
(m = DIV m n * n + r m n & < (r m n) n))" |
689 |
by (import hollight DEF_MOD) |
|
690 |
||
| 17652 | 691 |
lemma DIVISION: "ALL (m::nat) n::nat. n ~= 0 --> m = DIV m n * n + MOD m n & < (MOD m n) n" |
| 17644 | 692 |
by (import hollight DIVISION) |
693 |
||
694 |
lemma DIVMOD_UNIQ_LEMMA: "ALL (m::nat) (n::nat) (q1::nat) (r1::nat) (q2::nat) r2::nat. |
|
695 |
(m = q1 * n + r1 & < r1 n) & m = q2 * n + r2 & < r2 n --> |
|
696 |
q1 = q2 & r1 = r2" |
|
697 |
by (import hollight DIVMOD_UNIQ_LEMMA) |
|
698 |
||
699 |
lemma DIVMOD_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat. |
|
700 |
m = q * n + r & < r n --> DIV m n = q & MOD m n = r" |
|
701 |
by (import hollight DIVMOD_UNIQ) |
|
702 |
||
703 |
lemma MOD_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> MOD m n = r" |
|
704 |
by (import hollight MOD_UNIQ) |
|
705 |
||
706 |
lemma DIV_UNIQ: "ALL (m::nat) (n::nat) (q::nat) r::nat. m = q * n + r & < r n --> DIV m n = q" |
|
707 |
by (import hollight DIV_UNIQ) |
|
708 |
||
| 17652 | 709 |
lemma MOD_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> MOD (x * xa) x = 0" |
| 17644 | 710 |
by (import hollight MOD_MULT) |
711 |
||
| 17652 | 712 |
lemma DIV_MULT: "ALL (x::nat) xa::nat. x ~= 0 --> DIV (x * xa) x = xa" |
| 17644 | 713 |
by (import hollight DIV_MULT) |
714 |
||
| 17652 | 715 |
lemma DIV_DIV: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> DIV (DIV m n) p = DIV m (n * p)" |
| 17644 | 716 |
by (import hollight DIV_DIV) |
717 |
||
718 |
lemma MOD_LT: "ALL (m::nat) n::nat. < m n --> MOD m n = m" |
|
719 |
by (import hollight MOD_LT) |
|
720 |
||
721 |
lemma MOD_EQ: "ALL (m::nat) (n::nat) (p::nat) q::nat. m = n + q * p --> MOD m p = MOD n p" |
|
722 |
by (import hollight MOD_EQ) |
|
723 |
||
724 |
lemma DIV_MOD: "ALL (m::nat) (n::nat) p::nat. |
|
| 17652 | 725 |
n * p ~= 0 --> MOD (DIV m n) p = DIV (MOD m (n * p)) n" |
| 17644 | 726 |
by (import hollight DIV_MOD) |
727 |
||
| 17652 | 728 |
lemma DIV_1: "ALL n::nat. DIV n (NUMERAL_BIT1 0) = n" |
| 17644 | 729 |
by (import hollight DIV_1) |
730 |
||
| 17652 | 731 |
lemma EXP_LT_0: "ALL (x::nat) xa::nat. < 0 (EXP xa x) = (xa ~= 0 | x = 0)" |
| 17644 | 732 |
by (import hollight EXP_LT_0) |
733 |
||
| 17652 | 734 |
lemma DIV_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (DIV m n) m" |
| 17644 | 735 |
by (import hollight DIV_LE) |
736 |
||
737 |
lemma DIV_MUL_LE: "ALL (m::nat) n::nat. <= (n * DIV m n) m" |
|
738 |
by (import hollight DIV_MUL_LE) |
|
739 |
||
| 17652 | 740 |
lemma DIV_0: "ALL n::nat. n ~= 0 --> DIV 0 n = 0" |
| 17644 | 741 |
by (import hollight DIV_0) |
742 |
||
| 17652 | 743 |
lemma MOD_0: "ALL n::nat. n ~= 0 --> MOD 0 n = 0" |
| 17644 | 744 |
by (import hollight MOD_0) |
745 |
||
| 17652 | 746 |
lemma DIV_LT: "ALL (m::nat) n::nat. < m n --> DIV m n = 0" |
| 17644 | 747 |
by (import hollight DIV_LT) |
748 |
||
| 17652 | 749 |
lemma MOD_MOD: "ALL (m::nat) (n::nat) p::nat. n * p ~= 0 --> MOD (MOD m (n * p)) n = MOD m n" |
| 17644 | 750 |
by (import hollight MOD_MOD) |
751 |
||
| 17652 | 752 |
lemma MOD_MOD_REFL: "ALL (m::nat) n::nat. n ~= 0 --> MOD (MOD m n) n = MOD m n" |
| 17644 | 753 |
by (import hollight MOD_MOD_REFL) |
754 |
||
755 |
lemma DIV_MULT2: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 756 |
x * xb ~= 0 --> DIV (x * xa) (x * xb) = DIV xa xb" |
| 17644 | 757 |
by (import hollight DIV_MULT2) |
758 |
||
759 |
lemma MOD_MULT2: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 760 |
x * xb ~= 0 --> MOD (x * xa) (x * xb) = x * MOD xa xb" |
| 17644 | 761 |
by (import hollight MOD_MULT2) |
762 |
||
| 17652 | 763 |
lemma MOD_1: "ALL n::nat. MOD n (NUMERAL_BIT1 0) = 0" |
| 17644 | 764 |
by (import hollight MOD_1) |
765 |
||
766 |
lemma MOD_EXISTS: "ALL (m::nat) n::nat. |
|
| 17652 | 767 |
(EX q::nat. m = n * q) = COND (n = 0) (m = 0) (MOD m n = 0)" |
| 17644 | 768 |
by (import hollight MOD_EXISTS) |
769 |
||
770 |
lemma LT_EXP: "ALL (x::nat) (m::nat) n::nat. |
|
771 |
< (EXP x m) (EXP x n) = |
|
| 17652 | 772 |
(<= (NUMERAL_BIT0 (NUMERAL_BIT1 0)) x & < m n | x = 0 & m ~= 0 & n = 0)" |
| 17644 | 773 |
by (import hollight LT_EXP) |
774 |
||
775 |
lemma LE_EXP: "ALL (x::nat) (m::nat) n::nat. |
|
776 |
<= (EXP x m) (EXP x n) = |
|
| 17652 | 777 |
COND (x = 0) (m = 0 --> n = 0) (x = NUMERAL_BIT1 0 | <= m n)" |
| 17644 | 778 |
by (import hollight LE_EXP) |
779 |
||
| 17652 | 780 |
lemma DIV_MONO: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= m n --> <= (DIV m p) (DIV n p)" |
| 17644 | 781 |
by (import hollight DIV_MONO) |
782 |
||
783 |
lemma DIV_MONO_LT: "ALL (m::nat) (n::nat) p::nat. |
|
| 17652 | 784 |
p ~= 0 & <= (m + p) n --> < (DIV m p) (DIV n p)" |
| 17644 | 785 |
by (import hollight DIV_MONO_LT) |
786 |
||
| 17652 | 787 |
lemma LE_LDIV: "ALL (a::nat) (b::nat) n::nat. a ~= 0 & <= b (a * n) --> <= (DIV b a) n" |
| 17644 | 788 |
by (import hollight LE_LDIV) |
789 |
||
| 17652 | 790 |
lemma LE_RDIV_EQ: "ALL (a::nat) (b::nat) n::nat. a ~= 0 --> <= n (DIV b a) = <= (a * n) b" |
| 17644 | 791 |
by (import hollight LE_RDIV_EQ) |
792 |
||
793 |
lemma LE_LDIV_EQ: "ALL (a::nat) (b::nat) n::nat. |
|
| 17652 | 794 |
a ~= 0 --> <= (DIV b a) n = < b (a * (n + NUMERAL_BIT1 0))" |
| 17644 | 795 |
by (import hollight LE_LDIV_EQ) |
796 |
||
| 17652 | 797 |
lemma DIV_EQ_0: "ALL (m::nat) n::nat. n ~= 0 --> (DIV m n = 0) = < m n" |
| 17644 | 798 |
by (import hollight DIV_EQ_0) |
799 |
||
| 17652 | 800 |
lemma MOD_EQ_0: "ALL (m::nat) n::nat. n ~= 0 --> (MOD m n = 0) = (EX q::nat. m = q * n)" |
| 17644 | 801 |
by (import hollight MOD_EQ_0) |
802 |
||
| 17652 | 803 |
lemma EVEN_MOD: "ALL n::nat. EVEN n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = 0)" |
| 17644 | 804 |
by (import hollight EVEN_MOD) |
805 |
||
| 17652 | 806 |
lemma ODD_MOD: "ALL n::nat. ODD n = (MOD n (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = NUMERAL_BIT1 0)" |
| 17644 | 807 |
by (import hollight ODD_MOD) |
808 |
||
| 17652 | 809 |
lemma MOD_MULT_RMOD: "ALL (m::nat) (n::nat) p::nat. n ~= 0 --> MOD (m * MOD p n) n = MOD (m * p) n" |
| 17644 | 810 |
by (import hollight MOD_MULT_RMOD) |
811 |
||
812 |
lemma MOD_MULT_LMOD: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 813 |
xa ~= 0 --> MOD (MOD x xa * xb) xa = MOD (x * xb) xa" |
| 17644 | 814 |
by (import hollight MOD_MULT_LMOD) |
815 |
||
816 |
lemma MOD_MULT_MOD2: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 817 |
xa ~= 0 --> MOD (MOD x xa * MOD xb xa) xa = MOD (x * xb) xa" |
| 17644 | 818 |
by (import hollight MOD_MULT_MOD2) |
819 |
||
820 |
lemma MOD_EXP_MOD: "ALL (m::nat) (n::nat) p::nat. |
|
| 17652 | 821 |
n ~= 0 --> MOD (EXP (MOD m n) p) n = MOD (EXP m p) n" |
| 17644 | 822 |
by (import hollight MOD_EXP_MOD) |
823 |
||
824 |
lemma MOD_MULT_ADD: "ALL (m::nat) (n::nat) p::nat. MOD (m * n + p) n = MOD p n" |
|
825 |
by (import hollight MOD_MULT_ADD) |
|
826 |
||
827 |
lemma MOD_ADD_MOD: "ALL (a::nat) (b::nat) n::nat. |
|
| 17652 | 828 |
n ~= 0 --> MOD (MOD a n + MOD b n) n = MOD (a + b) n" |
| 17644 | 829 |
by (import hollight MOD_ADD_MOD) |
830 |
||
831 |
lemma DIV_ADD_MOD: "ALL (a::nat) (b::nat) n::nat. |
|
| 17652 | 832 |
n ~= 0 --> |
| 17644 | 833 |
(MOD (a + b) n = MOD a n + MOD b n) = (DIV (a + b) n = DIV a n + DIV b n)" |
834 |
by (import hollight DIV_ADD_MOD) |
|
835 |
||
| 17652 | 836 |
lemma DIV_REFL: "ALL n::nat. n ~= 0 --> DIV n n = NUMERAL_BIT1 0" |
| 17644 | 837 |
by (import hollight DIV_REFL) |
838 |
||
| 17652 | 839 |
lemma MOD_LE: "ALL (m::nat) n::nat. n ~= 0 --> <= (MOD m n) m" |
| 17644 | 840 |
by (import hollight MOD_LE) |
841 |
||
| 17652 | 842 |
lemma DIV_MONO2: "ALL (m::nat) (n::nat) p::nat. p ~= 0 & <= p m --> <= (DIV n m) (DIV n p)" |
| 17644 | 843 |
by (import hollight DIV_MONO2) |
844 |
||
845 |
lemma DIV_LE_EXCLUSION: "ALL (a::nat) (b::nat) (c::nat) d::nat. |
|
| 17652 | 846 |
b ~= 0 & < (b * c) ((a + NUMERAL_BIT1 0) * d) --> <= (DIV c d) (DIV a b)" |
| 17644 | 847 |
by (import hollight DIV_LE_EXCLUSION) |
848 |
||
| 17652 | 849 |
lemma DIV_EQ_EXCLUSION: "< ((b::nat) * (c::nat)) (((a::nat) + NUMERAL_BIT1 0) * (d::nat)) & |
850 |
< (a * d) ((c + NUMERAL_BIT1 0) * b) --> |
|
| 17644 | 851 |
DIV a b = DIV c d" |
852 |
by (import hollight DIV_EQ_EXCLUSION) |
|
853 |
||
854 |
lemma SUB_ELIM_THM: "(P::nat => bool) ((a::nat) - (b::nat)) = |
|
| 17652 | 855 |
(ALL x::nat. (b = a + x --> P 0) & (a = b + x --> P x))" |
| 17644 | 856 |
by (import hollight SUB_ELIM_THM) |
857 |
||
858 |
lemma PRE_ELIM_THM: "(P::nat => bool) (Pred (n::nat)) = |
|
| 17652 | 859 |
(ALL m::nat. (n = 0 --> P 0) & (n = Suc m --> P m))" |
| 17644 | 860 |
by (import hollight PRE_ELIM_THM) |
861 |
||
862 |
lemma DIVMOD_ELIM_THM: "(P::nat => nat => bool) (DIV (m::nat) (n::nat)) (MOD m n) = |
|
| 17652 | 863 |
(n = 0 & P 0 0 | |
864 |
n ~= 0 & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))" |
|
| 17644 | 865 |
by (import hollight DIVMOD_ELIM_THM) |
866 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
867 |
definition eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" where
|
| 17644 | 868 |
"eqeq == |
869 |
%(u::'q_9910::type) (ua::'q_9909::type) |
|
870 |
ub::'q_9910::type => 'q_9909::type => bool. ub u ua" |
|
871 |
||
872 |
lemma DEF__equal__equal_: "eqeq = |
|
873 |
(%(u::'q_9910::type) (ua::'q_9909::type) |
|
874 |
ub::'q_9910::type => 'q_9909::type => bool. ub u ua)" |
|
875 |
by (import hollight DEF__equal__equal_) |
|
876 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
877 |
definition mod_nat :: "nat => nat => nat => bool" where |
| 17644 | 878 |
"mod_nat == |
879 |
%(u::nat) (ua::nat) ub::nat. EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2" |
|
880 |
||
881 |
lemma DEF_mod_nat: "mod_nat = |
|
882 |
(%(u::nat) (ua::nat) ub::nat. |
|
883 |
EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2)" |
|
884 |
by (import hollight DEF_mod_nat) |
|
885 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
886 |
definition minimal :: "(nat => bool) => nat" where |
| 17644 | 887 |
"minimal == %u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m)" |
888 |
||
889 |
lemma DEF_minimal: "minimal = |
|
890 |
(%u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m))" |
|
891 |
by (import hollight DEF_minimal) |
|
892 |
||
893 |
lemma MINIMAL: "ALL P::nat => bool. |
|
894 |
Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))" |
|
895 |
by (import hollight MINIMAL) |
|
896 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
897 |
definition WF :: "('A => 'A => bool) => bool" where
|
| 17644 | 898 |
"WF == |
899 |
%u::'A::type => 'A::type => bool. |
|
900 |
ALL P::'A::type => bool. |
|
901 |
Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y))" |
|
902 |
||
903 |
lemma DEF_WF: "WF = |
|
904 |
(%u::'A::type => 'A::type => bool. |
|
905 |
ALL P::'A::type => bool. |
|
906 |
Ex P --> (EX x::'A::type. P x & (ALL y::'A::type. u y x --> ~ P y)))" |
|
907 |
by (import hollight DEF_WF) |
|
908 |
||
| 19093 | 909 |
lemma WF_EQ: "WF (u_353::'A::type => 'A::type => bool) = |
| 17644 | 910 |
(ALL P::'A::type => bool. |
| 19093 | 911 |
Ex P = (EX x::'A::type. P x & (ALL y::'A::type. u_353 y x --> ~ P y)))" |
| 17644 | 912 |
by (import hollight WF_EQ) |
913 |
||
| 19093 | 914 |
lemma WF_IND: "WF (u_353::'A::type => 'A::type => bool) = |
| 17644 | 915 |
(ALL P::'A::type => bool. |
| 19093 | 916 |
(ALL x::'A::type. (ALL y::'A::type. u_353 y x --> P y) --> P x) --> |
| 17644 | 917 |
All P)" |
918 |
by (import hollight WF_IND) |
|
919 |
||
| 19093 | 920 |
lemma WF_DCHAIN: "WF (u_353::'A::type => 'A::type => bool) = |
921 |
(~ (EX s::nat => 'A::type. ALL n::nat. u_353 (s (Suc n)) (s n)))" |
|
| 17644 | 922 |
by (import hollight WF_DCHAIN) |
923 |
||
| 19093 | 924 |
lemma WF_UREC: "WF (u_353::'A::type => 'A::type => bool) --> |
| 17644 | 925 |
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
|
926 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
|
| 19093 | 927 |
(ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> |
| 17644 | 928 |
(ALL (f::'A::type => 'B::type) g::'A::type => 'B::type. |
929 |
(ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) --> |
|
930 |
f = g))" |
|
931 |
by (import hollight WF_UREC) |
|
932 |
||
933 |
lemma WF_UREC_WF: "(ALL H::('A::type => bool) => 'A::type => bool.
|
|
934 |
(ALL (f::'A::type => bool) (g::'A::type => bool) x::'A::type. |
|
935 |
(ALL z::'A::type. |
|
| 19093 | 936 |
(u_353::'A::type => 'A::type => bool) z x --> f z = g z) --> |
| 17644 | 937 |
H f x = H g x) --> |
938 |
(ALL (f::'A::type => bool) g::'A::type => bool. |
|
939 |
(ALL x::'A::type. f x = H f x) & (ALL x::'A::type. g x = H g x) --> |
|
940 |
f = g)) --> |
|
| 19093 | 941 |
WF u_353" |
| 17644 | 942 |
by (import hollight WF_UREC_WF) |
943 |
||
| 19093 | 944 |
lemma WF_REC_INVARIANT: "WF (u_353::'A::type => 'A::type => bool) --> |
| 17644 | 945 |
(ALL (H::('A::type => 'B::type) => 'A::type => 'B::type)
|
946 |
S::'A::type => 'B::type => bool. |
|
947 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
|
| 19093 | 948 |
(ALL z::'A::type. u_353 z x --> f z = g z & S z (f z)) --> |
| 17644 | 949 |
H f x = H g x & S x (H f x)) --> |
950 |
(EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" |
|
951 |
by (import hollight WF_REC_INVARIANT) |
|
952 |
||
| 19093 | 953 |
lemma WF_REC: "WF (u_353::'A::type => 'A::type => bool) --> |
| 17644 | 954 |
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
|
955 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
|
| 19093 | 956 |
(ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> |
| 17644 | 957 |
(EX f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" |
958 |
by (import hollight WF_REC) |
|
959 |
||
960 |
lemma WF_REC_WF: "(ALL H::('A::type => nat) => 'A::type => nat.
|
|
961 |
(ALL (f::'A::type => nat) (g::'A::type => nat) x::'A::type. |
|
962 |
(ALL z::'A::type. |
|
| 19093 | 963 |
(u_353::'A::type => 'A::type => bool) z x --> f z = g z) --> |
| 17644 | 964 |
H f x = H g x) --> |
965 |
(EX f::'A::type => nat. ALL x::'A::type. f x = H f x)) --> |
|
| 19093 | 966 |
WF u_353" |
| 17644 | 967 |
by (import hollight WF_REC_WF) |
968 |
||
| 19093 | 969 |
lemma WF_EREC: "WF (u_353::'A::type => 'A::type => bool) --> |
| 17644 | 970 |
(ALL H::('A::type => 'B::type) => 'A::type => 'B::type.
|
971 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
|
| 19093 | 972 |
(ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) --> |
| 17644 | 973 |
(EX! f::'A::type => 'B::type. ALL x::'A::type. f x = H f x))" |
974 |
by (import hollight WF_EREC) |
|
975 |
||
976 |
lemma WF_SUBSET: "(ALL (x::'A::type) y::'A::type. |
|
| 19093 | 977 |
(u_353::'A::type => 'A::type => bool) x y --> |
978 |
(u_472::'A::type => 'A::type => bool) x y) & |
|
979 |
WF u_472 --> |
|
980 |
WF u_353" |
|
| 17644 | 981 |
by (import hollight WF_SUBSET) |
982 |
||
983 |
lemma WF_MEASURE_GEN: "ALL m::'A::type => 'B::type. |
|
| 19093 | 984 |
WF (u_353::'B::type => 'B::type => bool) --> |
985 |
WF (%(x::'A::type) x'::'A::type. u_353 (m x) (m x'))" |
|
| 17644 | 986 |
by (import hollight WF_MEASURE_GEN) |
987 |
||
988 |
lemma WF_LEX_DEPENDENT: "ALL (R::'A::type => 'A::type => bool) |
|
989 |
S::'A::type => 'B::type => 'B::type => bool. |
|
990 |
WF R & (ALL x::'A::type. WF (S x)) --> |
|
991 |
WF (GABS |
|
992 |
(%f::'A::type * 'B::type => 'A::type * 'B::type => bool. |
|
993 |
ALL (r1::'A::type) s1::'B::type. |
|
994 |
GEQ (f (r1, s1)) |
|
995 |
(GABS |
|
996 |
(%f::'A::type * 'B::type => bool. |
|
997 |
ALL (r2::'A::type) s2::'B::type. |
|
998 |
GEQ (f (r2, s2)) |
|
999 |
(R r1 r2 | r1 = r2 & S r1 s1 s2)))))" |
|
1000 |
by (import hollight WF_LEX_DEPENDENT) |
|
1001 |
||
1002 |
lemma WF_LEX: "ALL (x::'A::type => 'A::type => bool) xa::'B::type => 'B::type => bool. |
|
1003 |
WF x & WF xa --> |
|
1004 |
WF (GABS |
|
1005 |
(%f::'A::type * 'B::type => 'A::type * 'B::type => bool. |
|
1006 |
ALL (r1::'A::type) s1::'B::type. |
|
1007 |
GEQ (f (r1, s1)) |
|
1008 |
(GABS |
|
1009 |
(%f::'A::type * 'B::type => bool. |
|
1010 |
ALL (r2::'A::type) s2::'B::type. |
|
1011 |
GEQ (f (r2, s2)) (x r1 r2 | r1 = r2 & xa s1 s2)))))" |
|
1012 |
by (import hollight WF_LEX) |
|
1013 |
||
| 19093 | 1014 |
lemma WF_POINTWISE: "WF (u_353::'A::type => 'A::type => bool) & |
1015 |
WF (u_472::'B::type => 'B::type => bool) --> |
|
| 17644 | 1016 |
WF (GABS |
1017 |
(%f::'A::type * 'B::type => 'A::type * 'B::type => bool. |
|
1018 |
ALL (x1::'A::type) y1::'B::type. |
|
1019 |
GEQ (f (x1, y1)) |
|
1020 |
(GABS |
|
1021 |
(%f::'A::type * 'B::type => bool. |
|
1022 |
ALL (x2::'A::type) y2::'B::type. |
|
| 19093 | 1023 |
GEQ (f (x2, y2)) (u_353 x1 x2 & u_472 y1 y2)))))" |
| 17644 | 1024 |
by (import hollight WF_POINTWISE) |
1025 |
||
1026 |
lemma WF_num: "WF <" |
|
1027 |
by (import hollight WF_num) |
|
1028 |
||
1029 |
lemma WF_REC_num: "ALL H::(nat => 'A::type) => nat => 'A::type. |
|
1030 |
(ALL (f::nat => 'A::type) (g::nat => 'A::type) x::nat. |
|
1031 |
(ALL z::nat. < z x --> f z = g z) --> H f x = H g x) --> |
|
1032 |
(EX f::nat => 'A::type. ALL x::nat. f x = H f x)" |
|
1033 |
by (import hollight WF_REC_num) |
|
1034 |
||
1035 |
consts |
|
| 17652 | 1036 |
measure :: "('q_11107 => nat) => 'q_11107 => 'q_11107 => bool"
|
| 17644 | 1037 |
|
1038 |
defs |
|
1039 |
measure_def: "hollight.measure == |
|
1040 |
%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type. |
|
1041 |
< (u x) (u y)" |
|
1042 |
||
1043 |
lemma DEF_measure: "hollight.measure = |
|
1044 |
(%(u::'q_11107::type => nat) (x::'q_11107::type) y::'q_11107::type. |
|
1045 |
< (u x) (u y))" |
|
1046 |
by (import hollight DEF_measure) |
|
1047 |
||
1048 |
lemma WF_MEASURE: "ALL m::'A::type => nat. WF (hollight.measure m)" |
|
1049 |
by (import hollight WF_MEASURE) |
|
1050 |
||
1051 |
lemma MEASURE_LE: "(ALL x::'q_11137::type. |
|
1052 |
hollight.measure (m::'q_11137::type => nat) x (a::'q_11137::type) --> |
|
1053 |
hollight.measure m x (b::'q_11137::type)) = |
|
1054 |
<= (m a) (m b)" |
|
1055 |
by (import hollight MEASURE_LE) |
|
1056 |
||
| 19093 | 1057 |
lemma WF_REFL: "ALL x::'A::type. WF (u_353::'A::type => 'A::type => bool) --> ~ u_353 x x" |
| 17644 | 1058 |
by (import hollight WF_REFL) |
1059 |
||
1060 |
lemma WF_FALSE: "WF (%(x::'A::type) y::'A::type. False)" |
|
1061 |
by (import hollight WF_FALSE) |
|
1062 |
||
1063 |
lemma WF_REC_TAIL: "ALL (P::'A::type => bool) (g::'A::type => 'A::type) h::'A::type => 'B::type. |
|
1064 |
EX f::'A::type => 'B::type. |
|
1065 |
ALL x::'A::type. f x = COND (P x) (f (g x)) (h x)" |
|
1066 |
by (import hollight WF_REC_TAIL) |
|
1067 |
||
1068 |
lemma WF_REC_TAIL_GENERAL: "ALL (P::('A::type => 'B::type) => 'A::type => bool)
|
|
1069 |
(G::('A::type => 'B::type) => 'A::type => 'A::type)
|
|
1070 |
H::('A::type => 'B::type) => 'A::type => 'B::type.
|
|
| 19093 | 1071 |
WF (u_353::'A::type => 'A::type => bool) & |
| 17644 | 1072 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
| 19093 | 1073 |
(ALL z::'A::type. u_353 z x --> f z = g z) --> |
| 17644 | 1074 |
P f x = P g x & G f x = G g x & H f x = H g x) & |
1075 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) x::'A::type. |
|
| 19093 | 1076 |
(ALL z::'A::type. u_353 z x --> f z = g z) --> H f x = H g x) & |
| 17644 | 1077 |
(ALL (f::'A::type => 'B::type) (x::'A::type) y::'A::type. |
| 19093 | 1078 |
P f x & u_353 y (G f x) --> u_353 y x) --> |
| 17644 | 1079 |
(EX f::'A::type => 'B::type. |
1080 |
ALL x::'A::type. f x = COND (P f x) (f (G f x)) (H f x))" |
|
1081 |
by (import hollight WF_REC_TAIL_GENERAL) |
|
1082 |
||
| 17652 | 1083 |
lemma ARITH_ZERO: "(op &::bool => bool => bool) ((op =::nat => nat => bool) (0::nat) (0::nat)) |
1084 |
((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) (0::nat)) (0::nat))" |
|
| 17644 | 1085 |
by (import hollight ARITH_ZERO) |
1086 |
||
1087 |
lemma ARITH_SUC: "(ALL x::nat. Suc x = Suc x) & |
|
| 17652 | 1088 |
Suc 0 = NUMERAL_BIT1 0 & |
| 17644 | 1089 |
(ALL x::nat. Suc (NUMERAL_BIT0 x) = NUMERAL_BIT1 x) & |
1090 |
(ALL x::nat. Suc (NUMERAL_BIT1 x) = NUMERAL_BIT0 (Suc x))" |
|
1091 |
by (import hollight ARITH_SUC) |
|
1092 |
||
1093 |
lemma ARITH_PRE: "(ALL x::nat. Pred x = Pred x) & |
|
| 17652 | 1094 |
Pred 0 = 0 & |
| 17644 | 1095 |
(ALL x::nat. |
| 17652 | 1096 |
Pred (NUMERAL_BIT0 x) = COND (x = 0) 0 (NUMERAL_BIT1 (Pred x))) & |
| 17644 | 1097 |
(ALL x::nat. Pred (NUMERAL_BIT1 x) = NUMERAL_BIT0 x)" |
1098 |
by (import hollight ARITH_PRE) |
|
1099 |
||
| 17652 | 1100 |
lemma ARITH_ADD: "(op &::bool => bool => bool) |
1101 |
((All::(nat => bool) => bool) |
|
1102 |
(%x::nat. |
|
1103 |
(All::(nat => bool) => bool) |
|
1104 |
(%xa::nat. |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1105 |
(op =::nat => nat => bool) ((plus::nat => nat => nat) x xa) |
|
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1106 |
((plus::nat => nat => nat) x xa)))) |
| 17652 | 1107 |
((op &::bool => bool => bool) |
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1108 |
((op =::nat => nat => bool) ((plus::nat => nat => nat) (0::nat) (0::nat)) |
| 17652 | 1109 |
(0::nat)) |
1110 |
((op &::bool => bool => bool) |
|
1111 |
((All::(nat => bool) => bool) |
|
1112 |
(%x::nat. |
|
1113 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1114 |
((plus::nat => nat => nat) (0::nat) |
| 17652 | 1115 |
((NUMERAL_BIT0::nat => nat) x)) |
1116 |
((NUMERAL_BIT0::nat => nat) x))) |
|
1117 |
((op &::bool => bool => bool) |
|
1118 |
((All::(nat => bool) => bool) |
|
1119 |
(%x::nat. |
|
1120 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1121 |
((plus::nat => nat => nat) (0::nat) |
| 17652 | 1122 |
((NUMERAL_BIT1::nat => nat) x)) |
1123 |
((NUMERAL_BIT1::nat => nat) x))) |
|
1124 |
((op &::bool => bool => bool) |
|
1125 |
((All::(nat => bool) => bool) |
|
1126 |
(%x::nat. |
|
1127 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1128 |
((plus::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) |
| 17652 | 1129 |
(0::nat)) |
1130 |
((NUMERAL_BIT0::nat => nat) x))) |
|
1131 |
((op &::bool => bool => bool) |
|
1132 |
((All::(nat => bool) => bool) |
|
1133 |
(%x::nat. |
|
1134 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1135 |
((plus::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) |
| 17652 | 1136 |
(0::nat)) |
1137 |
((NUMERAL_BIT1::nat => nat) x))) |
|
1138 |
((op &::bool => bool => bool) |
|
1139 |
((All::(nat => bool) => bool) |
|
1140 |
(%x::nat. |
|
1141 |
(All::(nat => bool) => bool) |
|
1142 |
(%xa::nat. |
|
1143 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1144 |
((plus::nat => nat => nat) |
| 17652 | 1145 |
((NUMERAL_BIT0::nat => nat) x) |
1146 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
1147 |
((NUMERAL_BIT0::nat => nat) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1148 |
((plus::nat => nat => nat) x xa))))) |
| 17652 | 1149 |
((op &::bool => bool => bool) |
1150 |
((All::(nat => bool) => bool) |
|
1151 |
(%x::nat. |
|
1152 |
(All::(nat => bool) => bool) |
|
1153 |
(%xa::nat. |
|
1154 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1155 |
((plus::nat => nat => nat) |
| 17652 | 1156 |
((NUMERAL_BIT0::nat => nat) x) |
1157 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
1158 |
((NUMERAL_BIT1::nat => nat) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1159 |
((plus::nat => nat => nat) x xa))))) |
| 17652 | 1160 |
((op &::bool => bool => bool) |
1161 |
((All::(nat => bool) => bool) |
|
1162 |
(%x::nat. |
|
1163 |
(All::(nat => bool) => bool) |
|
1164 |
(%xa::nat. |
|
1165 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1166 |
((plus::nat => nat => nat) |
| 17652 | 1167 |
((NUMERAL_BIT1::nat => nat) x) |
1168 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
1169 |
((NUMERAL_BIT1::nat => nat) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1170 |
((plus::nat => nat => nat) x xa))))) |
| 17652 | 1171 |
((All::(nat => bool) => bool) |
1172 |
(%x::nat. |
|
1173 |
(All::(nat => bool) => bool) |
|
1174 |
(%xa::nat. |
|
1175 |
(op =::nat => nat => bool) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1176 |
((plus::nat => nat => nat) |
| 17652 | 1177 |
((NUMERAL_BIT1::nat => nat) x) |
1178 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
1179 |
((NUMERAL_BIT0::nat => nat) |
|
1180 |
((Suc::nat => nat) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1181 |
((plus::nat => nat => nat) x |
| 17652 | 1182 |
xa))))))))))))))" |
| 17644 | 1183 |
by (import hollight ARITH_ADD) |
1184 |
||
| 17652 | 1185 |
lemma ARITH_MULT: "(op &::bool => bool => bool) |
1186 |
((All::(nat => bool) => bool) |
|
1187 |
(%x::nat. |
|
1188 |
(All::(nat => bool) => bool) |
|
1189 |
(%xa::nat. |
|
1190 |
(op =::nat => nat => bool) ((op *::nat => nat => nat) x xa) |
|
1191 |
((op *::nat => nat => nat) x xa)))) |
|
1192 |
((op &::bool => bool => bool) |
|
1193 |
((op =::nat => nat => bool) ((op *::nat => nat => nat) (0::nat) (0::nat)) |
|
1194 |
(0::nat)) |
|
1195 |
((op &::bool => bool => bool) |
|
1196 |
((All::(nat => bool) => bool) |
|
1197 |
(%x::nat. |
|
1198 |
(op =::nat => nat => bool) |
|
1199 |
((op *::nat => nat => nat) (0::nat) |
|
1200 |
((NUMERAL_BIT0::nat => nat) x)) |
|
1201 |
(0::nat))) |
|
1202 |
((op &::bool => bool => bool) |
|
1203 |
((All::(nat => bool) => bool) |
|
1204 |
(%x::nat. |
|
1205 |
(op =::nat => nat => bool) |
|
1206 |
((op *::nat => nat => nat) (0::nat) |
|
1207 |
((NUMERAL_BIT1::nat => nat) x)) |
|
1208 |
(0::nat))) |
|
1209 |
((op &::bool => bool => bool) |
|
1210 |
((All::(nat => bool) => bool) |
|
1211 |
(%x::nat. |
|
1212 |
(op =::nat => nat => bool) |
|
1213 |
((op *::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) |
|
1214 |
(0::nat)) |
|
1215 |
(0::nat))) |
|
1216 |
((op &::bool => bool => bool) |
|
1217 |
((All::(nat => bool) => bool) |
|
1218 |
(%x::nat. |
|
1219 |
(op =::nat => nat => bool) |
|
1220 |
((op *::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) |
|
1221 |
(0::nat)) |
|
1222 |
(0::nat))) |
|
1223 |
((op &::bool => bool => bool) |
|
1224 |
((All::(nat => bool) => bool) |
|
1225 |
(%x::nat. |
|
1226 |
(All::(nat => bool) => bool) |
|
1227 |
(%xa::nat. |
|
1228 |
(op =::nat => nat => bool) |
|
1229 |
((op *::nat => nat => nat) |
|
1230 |
((NUMERAL_BIT0::nat => nat) x) |
|
1231 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
1232 |
((NUMERAL_BIT0::nat => nat) |
|
1233 |
((NUMERAL_BIT0::nat => nat) |
|
1234 |
((op *::nat => nat => nat) x xa)))))) |
|
1235 |
((op &::bool => bool => bool) |
|
1236 |
((All::(nat => bool) => bool) |
|
1237 |
(%x::nat. |
|
1238 |
(All::(nat => bool) => bool) |
|
1239 |
(%xa::nat. |
|
1240 |
(op =::nat => nat => bool) |
|
1241 |
((op *::nat => nat => nat) |
|
1242 |
((NUMERAL_BIT0::nat => nat) x) |
|
1243 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1244 |
((plus::nat => nat => nat) |
| 17652 | 1245 |
((NUMERAL_BIT0::nat => nat) x) |
1246 |
((NUMERAL_BIT0::nat => nat) |
|
1247 |
((NUMERAL_BIT0::nat => nat) |
|
1248 |
((op *::nat => nat => nat) x xa))))))) |
|
1249 |
((op &::bool => bool => bool) |
|
1250 |
((All::(nat => bool) => bool) |
|
1251 |
(%x::nat. |
|
1252 |
(All::(nat => bool) => bool) |
|
1253 |
(%xa::nat. |
|
1254 |
(op =::nat => nat => bool) |
|
1255 |
((op *::nat => nat => nat) |
|
1256 |
((NUMERAL_BIT1::nat => nat) x) |
|
1257 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1258 |
((plus::nat => nat => nat) |
| 17652 | 1259 |
((NUMERAL_BIT0::nat => nat) xa) |
1260 |
((NUMERAL_BIT0::nat => nat) |
|
1261 |
((NUMERAL_BIT0::nat => nat) |
|
1262 |
((op *::nat => nat => nat) x xa))))))) |
|
1263 |
((All::(nat => bool) => bool) |
|
1264 |
(%x::nat. |
|
1265 |
(All::(nat => bool) => bool) |
|
1266 |
(%xa::nat. |
|
1267 |
(op =::nat => nat => bool) |
|
1268 |
((op *::nat => nat => nat) |
|
1269 |
((NUMERAL_BIT1::nat => nat) x) |
|
1270 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1271 |
((plus::nat => nat => nat) |
| 17652 | 1272 |
((NUMERAL_BIT1::nat => nat) x) |
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
1273 |
((plus::nat => nat => nat) |
| 17652 | 1274 |
((NUMERAL_BIT0::nat => nat) xa) |
1275 |
((NUMERAL_BIT0::nat => nat) |
|
1276 |
((NUMERAL_BIT0::nat => nat) |
|
1277 |
((op *::nat => nat => nat) x |
|
1278 |
xa))))))))))))))))" |
|
| 17644 | 1279 |
by (import hollight ARITH_MULT) |
1280 |
||
1281 |
lemma ARITH_EXP: "(ALL (x::nat) xa::nat. EXP x xa = EXP x xa) & |
|
| 17652 | 1282 |
EXP 0 0 = NUMERAL_BIT1 0 & |
1283 |
(ALL m::nat. EXP (NUMERAL_BIT0 m) 0 = NUMERAL_BIT1 0) & |
|
1284 |
(ALL m::nat. EXP (NUMERAL_BIT1 m) 0 = NUMERAL_BIT1 0) & |
|
1285 |
(ALL n::nat. EXP 0 (NUMERAL_BIT0 n) = EXP 0 n * EXP 0 n) & |
|
| 17644 | 1286 |
(ALL (m::nat) n::nat. |
1287 |
EXP (NUMERAL_BIT0 m) (NUMERAL_BIT0 n) = |
|
1288 |
EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n) & |
|
1289 |
(ALL (m::nat) n::nat. |
|
1290 |
EXP (NUMERAL_BIT1 m) (NUMERAL_BIT0 n) = |
|
1291 |
EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n) & |
|
| 17652 | 1292 |
(ALL n::nat. EXP 0 (NUMERAL_BIT1 n) = 0) & |
| 17644 | 1293 |
(ALL (m::nat) n::nat. |
1294 |
EXP (NUMERAL_BIT0 m) (NUMERAL_BIT1 n) = |
|
1295 |
NUMERAL_BIT0 m * (EXP (NUMERAL_BIT0 m) n * EXP (NUMERAL_BIT0 m) n)) & |
|
1296 |
(ALL (m::nat) n::nat. |
|
1297 |
EXP (NUMERAL_BIT1 m) (NUMERAL_BIT1 n) = |
|
1298 |
NUMERAL_BIT1 m * (EXP (NUMERAL_BIT1 m) n * EXP (NUMERAL_BIT1 m) n))" |
|
1299 |
by (import hollight ARITH_EXP) |
|
1300 |
||
1301 |
lemma ARITH_EVEN: "(ALL x::nat. EVEN x = EVEN x) & |
|
| 17652 | 1302 |
EVEN 0 = True & |
| 17644 | 1303 |
(ALL x::nat. EVEN (NUMERAL_BIT0 x) = True) & |
1304 |
(ALL x::nat. EVEN (NUMERAL_BIT1 x) = False)" |
|
1305 |
by (import hollight ARITH_EVEN) |
|
1306 |
||
1307 |
lemma ARITH_ODD: "(ALL x::nat. ODD x = ODD x) & |
|
| 17652 | 1308 |
ODD 0 = False & |
| 17644 | 1309 |
(ALL x::nat. ODD (NUMERAL_BIT0 x) = False) & |
1310 |
(ALL x::nat. ODD (NUMERAL_BIT1 x) = True)" |
|
1311 |
by (import hollight ARITH_ODD) |
|
1312 |
||
1313 |
lemma ARITH_LE: "(ALL (x::nat) xa::nat. <= x xa = <= x xa) & |
|
| 17652 | 1314 |
<= 0 0 = True & |
1315 |
(ALL x::nat. <= (NUMERAL_BIT0 x) 0 = (x = 0)) & |
|
1316 |
(ALL x::nat. <= (NUMERAL_BIT1 x) 0 = False) & |
|
1317 |
(ALL x::nat. <= 0 (NUMERAL_BIT0 x) = True) & |
|
1318 |
(ALL x::nat. <= 0 (NUMERAL_BIT1 x) = True) & |
|
| 17644 | 1319 |
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = <= x xa) & |
1320 |
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) & |
|
1321 |
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) & |
|
1322 |
(ALL (x::nat) xa::nat. <= (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = <= x xa)" |
|
1323 |
by (import hollight ARITH_LE) |
|
1324 |
||
1325 |
lemma ARITH_LT: "(ALL (x::nat) xa::nat. < x xa = < x xa) & |
|
| 17652 | 1326 |
< 0 0 = False & |
1327 |
(ALL x::nat. < (NUMERAL_BIT0 x) 0 = False) & |
|
1328 |
(ALL x::nat. < (NUMERAL_BIT1 x) 0 = False) & |
|
1329 |
(ALL x::nat. < 0 (NUMERAL_BIT0 x) = < 0 x) & |
|
1330 |
(ALL x::nat. < 0 (NUMERAL_BIT1 x) = True) & |
|
| 17644 | 1331 |
(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT0 xa) = < x xa) & |
1332 |
(ALL (x::nat) xa::nat. < (NUMERAL_BIT0 x) (NUMERAL_BIT1 xa) = <= x xa) & |
|
1333 |
(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT0 xa) = < x xa) & |
|
1334 |
(ALL (x::nat) xa::nat. < (NUMERAL_BIT1 x) (NUMERAL_BIT1 xa) = < x xa)" |
|
1335 |
by (import hollight ARITH_LT) |
|
1336 |
||
| 17652 | 1337 |
lemma ARITH_EQ: "(op &::bool => bool => bool) |
1338 |
((All::(nat => bool) => bool) |
|
1339 |
(%x::nat. |
|
1340 |
(All::(nat => bool) => bool) |
|
1341 |
(%xa::nat. |
|
1342 |
(op =::bool => bool => bool) ((op =::nat => nat => bool) x xa) |
|
1343 |
((op =::nat => nat => bool) x xa)))) |
|
1344 |
((op &::bool => bool => bool) |
|
1345 |
((op =::bool => bool => bool) |
|
1346 |
((op =::nat => nat => bool) (0::nat) (0::nat)) (True::bool)) |
|
1347 |
((op &::bool => bool => bool) |
|
1348 |
((All::(nat => bool) => bool) |
|
1349 |
(%x::nat. |
|
1350 |
(op =::bool => bool => bool) |
|
1351 |
((op =::nat => nat => bool) ((NUMERAL_BIT0::nat => nat) x) |
|
1352 |
(0::nat)) |
|
1353 |
((op =::nat => nat => bool) x (0::nat)))) |
|
1354 |
((op &::bool => bool => bool) |
|
1355 |
((All::(nat => bool) => bool) |
|
1356 |
(%x::nat. |
|
1357 |
(op =::bool => bool => bool) |
|
1358 |
((op =::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) x) |
|
1359 |
(0::nat)) |
|
1360 |
(False::bool))) |
|
1361 |
((op &::bool => bool => bool) |
|
1362 |
((All::(nat => bool) => bool) |
|
1363 |
(%x::nat. |
|
1364 |
(op =::bool => bool => bool) |
|
1365 |
((op =::nat => nat => bool) (0::nat) |
|
1366 |
((NUMERAL_BIT0::nat => nat) x)) |
|
1367 |
((op =::nat => nat => bool) (0::nat) x))) |
|
1368 |
((op &::bool => bool => bool) |
|
1369 |
((All::(nat => bool) => bool) |
|
1370 |
(%x::nat. |
|
1371 |
(op =::bool => bool => bool) |
|
1372 |
((op =::nat => nat => bool) (0::nat) |
|
1373 |
((NUMERAL_BIT1::nat => nat) x)) |
|
1374 |
(False::bool))) |
|
1375 |
((op &::bool => bool => bool) |
|
1376 |
((All::(nat => bool) => bool) |
|
1377 |
(%x::nat. |
|
1378 |
(All::(nat => bool) => bool) |
|
1379 |
(%xa::nat. |
|
1380 |
(op =::bool => bool => bool) |
|
1381 |
((op =::nat => nat => bool) |
|
1382 |
((NUMERAL_BIT0::nat => nat) x) |
|
1383 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
1384 |
((op =::nat => nat => bool) x xa)))) |
|
1385 |
((op &::bool => bool => bool) |
|
1386 |
((All::(nat => bool) => bool) |
|
1387 |
(%x::nat. |
|
1388 |
(All::(nat => bool) => bool) |
|
1389 |
(%xa::nat. |
|
1390 |
(op =::bool => bool => bool) |
|
1391 |
((op =::nat => nat => bool) |
|
1392 |
((NUMERAL_BIT0::nat => nat) x) |
|
1393 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
1394 |
(False::bool)))) |
|
1395 |
((op &::bool => bool => bool) |
|
1396 |
((All::(nat => bool) => bool) |
|
1397 |
(%x::nat. |
|
1398 |
(All::(nat => bool) => bool) |
|
1399 |
(%xa::nat. |
|
1400 |
(op =::bool => bool => bool) |
|
1401 |
((op =::nat => nat => bool) |
|
1402 |
((NUMERAL_BIT1::nat => nat) x) |
|
1403 |
((NUMERAL_BIT0::nat => nat) xa)) |
|
1404 |
(False::bool)))) |
|
1405 |
((All::(nat => bool) => bool) |
|
1406 |
(%x::nat. |
|
1407 |
(All::(nat => bool) => bool) |
|
1408 |
(%xa::nat. |
|
1409 |
(op =::bool => bool => bool) |
|
1410 |
((op =::nat => nat => bool) |
|
1411 |
((NUMERAL_BIT1::nat => nat) x) |
|
1412 |
((NUMERAL_BIT1::nat => nat) xa)) |
|
1413 |
((op =::nat => nat => bool) x xa))))))))))))" |
|
| 17644 | 1414 |
by (import hollight ARITH_EQ) |
1415 |
||
| 17652 | 1416 |
lemma ARITH_SUB: "(op &::bool => bool => bool) |
1417 |
((All::(nat => bool) => bool) |
|
1418 |
(%x::nat. |
|
1419 |
(All::(nat => bool) => bool) |
|
1420 |
(%xa::nat. |
|
1421 |
(op =::nat => nat => bool) ((op -::nat => nat => nat) x xa) |
|
1422 |
((op -::nat => nat => nat) x xa)))) |
|
1423 |
((op &::bool => bool => bool) |
|
1424 |
((op =::nat => nat => bool) ((op -::nat => nat => nat) (0::nat) (0::nat)) |
|
1425 |
(0::nat)) |
|
1426 |
((op &::bool => bool => bool) |
|
1427 |
((All::(nat => bool) => bool) |
|
1428 |
(%x::nat. |
|
1429 |
(op =::nat => nat => bool) |
|
1430 |
((op -::nat => nat => nat) (0::nat) |
|
1431 |
((NUMERAL_BIT0::nat => nat) x)) |
|
1432 |
(0::nat))) |
|
1433 |
((op &::bool => bool => bool) |
|
1434 |
((All::(nat => bool) => bool) |
|
1435 |
(%x::nat. |
|
1436 |
(op =::nat => nat => bool) |
|
1437 |
((op -::nat => nat => nat) (0::nat) |
|
1438 |
((NUMERAL_BIT1::nat => nat) x)) |
|
1439 |
(0::nat))) |
|
1440 |
((op &::bool => bool => bool) |
|
1441 |
((All::(nat => bool) => bool) |
|
1442 |
(%x::nat. |
|
1443 |
(op =::nat => nat => bool) |
|
1444 |
((op -::nat => nat => nat) ((NUMERAL_BIT0::nat => nat) x) |
|
1445 |
(0::nat)) |
|
1446 |
((NUMERAL_BIT0::nat => nat) x))) |
|
1447 |
((op &::bool => bool => bool) |
|
1448 |
((All::(nat => bool) => bool) |
|
1449 |
(%x::nat. |
|
1450 |
(op =::nat => nat => bool) |
|
1451 |
((op -::nat => nat => nat) ((NUMERAL_BIT1::nat => nat) x) |
|
1452 |
(0::nat)) |
|
1453 |
((NUMERAL_BIT1::nat => nat) x))) |
|
1454 |
((op &::bool => bool => bool) |
|
1455 |
((All::(nat => bool) => bool) |
|
1456 |
(%m::nat. |
|
1457 |
(All::(nat => bool) => bool) |
|
1458 |
(%n::nat. |
|
1459 |
(op =::nat => nat => bool) |
|
1460 |
((op -::nat => nat => nat) |
|
1461 |
((NUMERAL_BIT0::nat => nat) m) |
|
1462 |
((NUMERAL_BIT0::nat => nat) n)) |
|
1463 |
((NUMERAL_BIT0::nat => nat) |
|
1464 |
((op -::nat => nat => nat) m n))))) |
|
1465 |
((op &::bool => bool => bool) |
|
1466 |
((All::(nat => bool) => bool) |
|
1467 |
(%m::nat. |
|
1468 |
(All::(nat => bool) => bool) |
|
1469 |
(%n::nat. |
|
1470 |
(op =::nat => nat => bool) |
|
1471 |
((op -::nat => nat => nat) |
|
1472 |
((NUMERAL_BIT0::nat => nat) m) |
|
1473 |
((NUMERAL_BIT1::nat => nat) n)) |
|
1474 |
((Pred::nat => nat) |
|
1475 |
((NUMERAL_BIT0::nat => nat) |
|
1476 |
((op -::nat => nat => nat) m n)))))) |
|
1477 |
((op &::bool => bool => bool) |
|
1478 |
((All::(nat => bool) => bool) |
|
1479 |
(%m::nat. |
|
1480 |
(All::(nat => bool) => bool) |
|
1481 |
(%n::nat. |
|
1482 |
(op =::nat => nat => bool) |
|
1483 |
((op -::nat => nat => nat) |
|
1484 |
((NUMERAL_BIT1::nat => nat) m) |
|
1485 |
((NUMERAL_BIT0::nat => nat) n)) |
|
1486 |
((COND::bool => nat => nat => nat) |
|
1487 |
((<=::nat => nat => bool) n m) |
|
1488 |
((NUMERAL_BIT1::nat => nat) |
|
1489 |
((op -::nat => nat => nat) m n)) |
|
1490 |
(0::nat))))) |
|
1491 |
((All::(nat => bool) => bool) |
|
1492 |
(%m::nat. |
|
1493 |
(All::(nat => bool) => bool) |
|
1494 |
(%n::nat. |
|
1495 |
(op =::nat => nat => bool) |
|
1496 |
((op -::nat => nat => nat) |
|
1497 |
((NUMERAL_BIT1::nat => nat) m) |
|
1498 |
((NUMERAL_BIT1::nat => nat) n)) |
|
1499 |
((NUMERAL_BIT0::nat => nat) |
|
1500 |
((op -::nat => nat => nat) m n)))))))))))))" |
|
| 17644 | 1501 |
by (import hollight ARITH_SUB) |
1502 |
||
1503 |
lemma right_th: "(s::nat) * NUMERAL_BIT1 (x::nat) = s + NUMERAL_BIT0 (s * x)" |
|
1504 |
by (import hollight right_th) |
|
1505 |
||
1506 |
lemma SEMIRING_PTHS: "(ALL (x::'A::type) (y::'A::type) z::'A::type. |
|
1507 |
(add::'A::type => 'A::type => 'A::type) x (add y z) = add (add x y) z) & |
|
1508 |
(ALL (x::'A::type) y::'A::type. add x y = add y x) & |
|
1509 |
(ALL x::'A::type. add (r0::'A::type) x = x) & |
|
1510 |
(ALL (x::'A::type) (y::'A::type) z::'A::type. |
|
1511 |
(mul::'A::type => 'A::type => 'A::type) x (mul y z) = mul (mul x y) z) & |
|
1512 |
(ALL (x::'A::type) y::'A::type. mul x y = mul y x) & |
|
1513 |
(ALL x::'A::type. mul (r1::'A::type) x = x) & |
|
1514 |
(ALL x::'A::type. mul r0 x = r0) & |
|
1515 |
(ALL (x::'A::type) (y::'A::type) z::'A::type. |
|
1516 |
mul x (add y z) = add (mul x y) (mul x z)) & |
|
| 17652 | 1517 |
(ALL x::'A::type. (pwr::'A::type => nat => 'A::type) x 0 = r1) & |
| 17644 | 1518 |
(ALL (x::'A::type) n::nat. pwr x (Suc n) = mul x (pwr x n)) --> |
1519 |
mul r1 (x::'A::type) = x & |
|
1520 |
add (mul (a::'A::type) (m::'A::type)) (mul (b::'A::type) m) = |
|
1521 |
mul (add a b) m & |
|
1522 |
add (mul a m) m = mul (add a r1) m & |
|
1523 |
add m (mul a m) = mul (add a r1) m & |
|
1524 |
add m m = mul (add r1 r1) m & |
|
1525 |
mul r0 m = r0 & |
|
1526 |
add r0 a = a & |
|
1527 |
add a r0 = a & |
|
1528 |
mul a b = mul b a & |
|
1529 |
mul (add a b) (c::'A::type) = add (mul a c) (mul b c) & |
|
1530 |
mul r0 a = r0 & |
|
1531 |
mul a r0 = r0 & |
|
1532 |
mul r1 a = a & |
|
1533 |
mul a r1 = a & |
|
1534 |
mul (mul (lx::'A::type) (ly::'A::type)) |
|
1535 |
(mul (rx::'A::type) (ry::'A::type)) = |
|
1536 |
mul (mul lx rx) (mul ly ry) & |
|
1537 |
mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry)) & |
|
1538 |
mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry) & |
|
1539 |
mul (mul lx ly) rx = mul (mul lx rx) ly & |
|
1540 |
mul (mul lx ly) rx = mul lx (mul ly rx) & |
|
1541 |
mul lx rx = mul rx lx & |
|
1542 |
mul lx (mul rx ry) = mul (mul lx rx) ry & |
|
1543 |
mul lx (mul rx ry) = mul rx (mul lx ry) & |
|
1544 |
add (add a b) (add c (d::'A::type)) = add (add a c) (add b d) & |
|
1545 |
add (add a b) c = add a (add b c) & |
|
1546 |
add a (add c d) = add c (add a d) & |
|
1547 |
add (add a b) c = add (add a c) b & |
|
1548 |
add a c = add c a & |
|
1549 |
add a (add c d) = add (add a c) d & |
|
1550 |
mul (pwr x (p::nat)) (pwr x (q::nat)) = pwr x (p + q) & |
|
1551 |
mul x (pwr x q) = pwr x (Suc q) & |
|
1552 |
mul (pwr x q) x = pwr x (Suc q) & |
|
| 17652 | 1553 |
mul x x = pwr x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) & |
| 17644 | 1554 |
pwr (mul x (y::'A::type)) q = mul (pwr x q) (pwr y q) & |
1555 |
pwr (pwr x p) q = pwr x (p * q) & |
|
| 17652 | 1556 |
pwr x 0 = r1 & |
1557 |
pwr x (NUMERAL_BIT1 0) = x & |
|
| 17644 | 1558 |
mul x (add y (z::'A::type)) = add (mul x y) (mul x z) & |
1559 |
pwr x (Suc q) = mul x (pwr x q)" |
|
1560 |
by (import hollight SEMIRING_PTHS) |
|
1561 |
||
1562 |
lemma sth: "(ALL (x::nat) (y::nat) z::nat. x + (y + z) = x + y + z) & |
|
1563 |
(ALL (x::nat) y::nat. x + y = y + x) & |
|
| 17652 | 1564 |
(ALL x::nat. 0 + x = x) & |
| 17644 | 1565 |
(ALL (x::nat) (y::nat) z::nat. x * (y * z) = x * y * z) & |
1566 |
(ALL (x::nat) y::nat. x * y = y * x) & |
|
| 17652 | 1567 |
(ALL x::nat. NUMERAL_BIT1 0 * x = x) & |
1568 |
(ALL x::nat. 0 * x = 0) & |
|
| 17644 | 1569 |
(ALL (x::nat) (xa::nat) xb::nat. x * (xa + xb) = x * xa + x * xb) & |
| 17652 | 1570 |
(ALL x::nat. EXP x 0 = NUMERAL_BIT1 0) & |
| 17644 | 1571 |
(ALL (x::nat) xa::nat. EXP x (Suc xa) = x * EXP x xa)" |
1572 |
by (import hollight sth) |
|
1573 |
||
1574 |
lemma NUM_INTEGRAL_LEMMA: "(w::nat) = (x::nat) + (d::nat) & (y::nat) = (z::nat) + (e::nat) --> |
|
1575 |
(w * y + x * z = w * z + x * y) = (w = x | y = z)" |
|
1576 |
by (import hollight NUM_INTEGRAL_LEMMA) |
|
1577 |
||
| 17652 | 1578 |
lemma NUM_INTEGRAL: "(ALL x::nat. 0 * x = 0) & |
| 17644 | 1579 |
(ALL (x::nat) (xa::nat) xb::nat. (x + xa = x + xb) = (xa = xb)) & |
1580 |
(ALL (w::nat) (x::nat) (y::nat) z::nat. |
|
1581 |
(w * y + x * z = w * z + x * y) = (w = x | y = z))" |
|
1582 |
by (import hollight NUM_INTEGRAL) |
|
1583 |
||
1584 |
lemma INJ_INVERSE2: "ALL P::'A::type => 'B::type => 'C::type. |
|
1585 |
(ALL (x1::'A::type) (y1::'B::type) (x2::'A::type) y2::'B::type. |
|
1586 |
(P x1 y1 = P x2 y2) = (x1 = x2 & y1 = y2)) --> |
|
1587 |
(EX (x::'C::type => 'A::type) Y::'C::type => 'B::type. |
|
1588 |
ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)" |
|
1589 |
by (import hollight INJ_INVERSE2) |
|
1590 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1591 |
definition NUMPAIR :: "nat => nat => nat" where |
| 17644 | 1592 |
"NUMPAIR == |
1593 |
%(u::nat) ua::nat. |
|
| 17652 | 1594 |
EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u * |
1595 |
(NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0)" |
|
| 17644 | 1596 |
|
1597 |
lemma DEF_NUMPAIR: "NUMPAIR = |
|
1598 |
(%(u::nat) ua::nat. |
|
| 17652 | 1599 |
EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u * |
1600 |
(NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua + NUMERAL_BIT1 0))" |
|
| 17644 | 1601 |
by (import hollight DEF_NUMPAIR) |
1602 |
||
1603 |
lemma NUMPAIR_INJ_LEMMA: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. |
|
1604 |
NUMPAIR x xa = NUMPAIR xb xc --> x = xb" |
|
1605 |
by (import hollight NUMPAIR_INJ_LEMMA) |
|
1606 |
||
1607 |
lemma NUMPAIR_INJ: "ALL (x1::nat) (y1::nat) (x2::nat) y2::nat. |
|
1608 |
(NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)" |
|
1609 |
by (import hollight NUMPAIR_INJ) |
|
1610 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1611 |
definition NUMFST :: "nat => nat" where |
| 17644 | 1612 |
"NUMFST == |
1613 |
SOME X::nat => nat. |
|
1614 |
EX Y::nat => nat. |
|
1615 |
ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y" |
|
1616 |
||
1617 |
lemma DEF_NUMFST: "NUMFST = |
|
1618 |
(SOME X::nat => nat. |
|
1619 |
EX Y::nat => nat. |
|
1620 |
ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" |
|
1621 |
by (import hollight DEF_NUMFST) |
|
1622 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1623 |
definition NUMSND :: "nat => nat" where |
| 17644 | 1624 |
"NUMSND == |
1625 |
SOME Y::nat => nat. |
|
1626 |
ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y" |
|
1627 |
||
1628 |
lemma DEF_NUMSND: "NUMSND = |
|
1629 |
(SOME Y::nat => nat. |
|
1630 |
ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" |
|
1631 |
by (import hollight DEF_NUMSND) |
|
1632 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1633 |
definition NUMSUM :: "bool => nat => nat" where |
| 17644 | 1634 |
"NUMSUM == |
1635 |
%(u::bool) ua::nat. |
|
| 17652 | 1636 |
COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)) |
1637 |
(NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)" |
|
| 17644 | 1638 |
|
1639 |
lemma DEF_NUMSUM: "NUMSUM = |
|
1640 |
(%(u::bool) ua::nat. |
|
| 17652 | 1641 |
COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)) |
1642 |
(NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua))" |
|
| 17644 | 1643 |
by (import hollight DEF_NUMSUM) |
1644 |
||
1645 |
lemma NUMSUM_INJ: "ALL (b1::bool) (x1::nat) (b2::bool) x2::nat. |
|
1646 |
(NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)" |
|
1647 |
by (import hollight NUMSUM_INJ) |
|
1648 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1649 |
definition NUMLEFT :: "nat => bool" where |
| 17644 | 1650 |
"NUMLEFT == |
1651 |
SOME X::nat => bool. |
|
1652 |
EX Y::nat => nat. |
|
1653 |
ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y" |
|
1654 |
||
1655 |
lemma DEF_NUMLEFT: "NUMLEFT = |
|
1656 |
(SOME X::nat => bool. |
|
1657 |
EX Y::nat => nat. |
|
1658 |
ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)" |
|
1659 |
by (import hollight DEF_NUMLEFT) |
|
1660 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1661 |
definition NUMRIGHT :: "nat => nat" where |
| 17644 | 1662 |
"NUMRIGHT == |
1663 |
SOME Y::nat => nat. |
|
1664 |
ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y" |
|
1665 |
||
1666 |
lemma DEF_NUMRIGHT: "NUMRIGHT = |
|
1667 |
(SOME Y::nat => nat. |
|
1668 |
ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)" |
|
1669 |
by (import hollight DEF_NUMRIGHT) |
|
1670 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1671 |
definition INJN :: "nat => nat => 'A => bool" where |
| 17644 | 1672 |
"INJN == %(u::nat) (n::nat) a::'A::type. n = u" |
1673 |
||
1674 |
lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. n = u)" |
|
1675 |
by (import hollight DEF_INJN) |
|
1676 |
||
1677 |
lemma INJN_INJ: "(All::(nat => bool) => bool) |
|
1678 |
(%n1::nat. |
|
1679 |
(All::(nat => bool) => bool) |
|
1680 |
(%n2::nat. |
|
1681 |
(op =::bool => bool => bool) |
|
1682 |
((op =::(nat => 'A::type => bool) |
|
1683 |
=> (nat => 'A::type => bool) => bool) |
|
1684 |
((INJN::nat => nat => 'A::type => bool) n1) |
|
1685 |
((INJN::nat => nat => 'A::type => bool) n2)) |
|
1686 |
((op =::nat => nat => bool) n1 n2)))" |
|
1687 |
by (import hollight INJN_INJ) |
|
1688 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1689 |
definition INJA :: "'A => nat => 'A => bool" where |
| 17644 | 1690 |
"INJA == %(u::'A::type) (n::nat) b::'A::type. b = u" |
1691 |
||
1692 |
lemma DEF_INJA: "INJA = (%(u::'A::type) (n::nat) b::'A::type. b = u)" |
|
1693 |
by (import hollight DEF_INJA) |
|
1694 |
||
1695 |
lemma INJA_INJ: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)" |
|
1696 |
by (import hollight INJA_INJ) |
|
1697 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1698 |
definition INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" where |
| 17644 | 1699 |
"INJF == %(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n)" |
1700 |
||
1701 |
lemma DEF_INJF: "INJF = |
|
1702 |
(%(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n))" |
|
1703 |
by (import hollight DEF_INJF) |
|
1704 |
||
1705 |
lemma INJF_INJ: "ALL (f1::nat => nat => 'A::type => bool) f2::nat => nat => 'A::type => bool. |
|
1706 |
(INJF f1 = INJF f2) = (f1 = f2)" |
|
1707 |
by (import hollight INJF_INJ) |
|
1708 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1709 |
definition INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where |
| 17644 | 1710 |
"INJP == |
1711 |
%(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat) |
|
1712 |
a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a)" |
|
1713 |
||
1714 |
lemma DEF_INJP: "INJP = |
|
1715 |
(%(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat) |
|
1716 |
a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a))" |
|
1717 |
by (import hollight DEF_INJP) |
|
1718 |
||
1719 |
lemma INJP_INJ: "ALL (f1::nat => 'A::type => bool) (f1'::nat => 'A::type => bool) |
|
1720 |
(f2::nat => 'A::type => bool) f2'::nat => 'A::type => bool. |
|
1721 |
(INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')" |
|
1722 |
by (import hollight INJP_INJ) |
|
1723 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1724 |
definition ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where |
| 17644 | 1725 |
"ZCONSTR == |
1726 |
%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool. |
|
1727 |
INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))" |
|
1728 |
||
1729 |
lemma DEF_ZCONSTR: "ZCONSTR = |
|
1730 |
(%(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool. |
|
1731 |
INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))" |
|
1732 |
by (import hollight DEF_ZCONSTR) |
|
1733 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1734 |
definition ZBOT :: "nat => 'A => bool" where |
| 17652 | 1735 |
"ZBOT == INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)" |
1736 |
||
1737 |
lemma DEF_ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)" |
|
| 17644 | 1738 |
by (import hollight DEF_ZBOT) |
1739 |
||
1740 |
lemma ZCONSTR_ZBOT: "ALL (x::nat) (xa::'A::type) xb::nat => nat => 'A::type => bool. |
|
1741 |
ZCONSTR x xa xb ~= ZBOT" |
|
1742 |
by (import hollight ZCONSTR_ZBOT) |
|
1743 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1744 |
definition ZRECSPACE :: "(nat => 'A => bool) => bool" where |
| 17644 | 1745 |
"ZRECSPACE == |
1746 |
%a::nat => 'A::type => bool. |
|
1747 |
ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. |
|
1748 |
(ALL a::nat => 'A::type => bool. |
|
1749 |
a = ZBOT | |
|
1750 |
(EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool. |
|
1751 |
a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) --> |
|
1752 |
ZRECSPACE' a) --> |
|
1753 |
ZRECSPACE' a" |
|
1754 |
||
1755 |
lemma DEF_ZRECSPACE: "ZRECSPACE = |
|
1756 |
(%a::nat => 'A::type => bool. |
|
1757 |
ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. |
|
1758 |
(ALL a::nat => 'A::type => bool. |
|
1759 |
a = ZBOT | |
|
1760 |
(EX (c::nat) (i::'A::type) r::nat => nat => 'A::type => bool. |
|
1761 |
a = ZCONSTR c i r & (ALL n::nat. ZRECSPACE' (r n))) --> |
|
1762 |
ZRECSPACE' a) --> |
|
1763 |
ZRECSPACE' a)" |
|
1764 |
by (import hollight DEF_ZRECSPACE) |
|
1765 |
||
1766 |
typedef (open) ('A) recspace = "(Collect::((nat => 'A::type => bool) => bool)
|
|
1767 |
=> (nat => 'A::type => bool) set) |
|
1768 |
(ZRECSPACE::(nat => 'A::type => bool) => bool)" morphisms "_dest_rec" "_mk_rec" |
|
1769 |
apply (rule light_ex_imp_nonempty[where t="ZBOT::nat => 'A::type => bool"]) |
|
1770 |
by (import hollight TYDEF_recspace) |
|
1771 |
||
1772 |
syntax |
|
1773 |
"_dest_rec" :: _ ("'_dest'_rec")
|
|
1774 |
||
1775 |
syntax |
|
1776 |
"_mk_rec" :: _ ("'_mk'_rec")
|
|
1777 |
||
1778 |
lemmas "TYDEF_recspace_@intern" = typedef_hol2hollight |
|
| 17652 | 1779 |
[where a="a :: 'A recspace" and r=r , |
| 17644 | 1780 |
OF type_definition_recspace] |
1781 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1782 |
definition BOTTOM :: "'A recspace" where |
| 17652 | 1783 |
"(op ==::'A::type recspace => 'A::type recspace => prop) |
1784 |
(BOTTOM::'A::type recspace) |
|
1785 |
((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) |
|
1786 |
(ZBOT::nat => 'A::type => bool))" |
|
1787 |
||
1788 |
lemma DEF_BOTTOM: "(op =::'A::type recspace => 'A::type recspace => bool) |
|
1789 |
(BOTTOM::'A::type recspace) |
|
1790 |
((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) |
|
1791 |
(ZBOT::nat => 'A::type => bool))" |
|
| 17644 | 1792 |
by (import hollight DEF_BOTTOM) |
1793 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1794 |
definition CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" where |
| 17652 | 1795 |
"(op ==::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
1796 |
=> (nat |
|
1797 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1798 |
=> prop) |
|
1799 |
(CONSTR::nat |
|
1800 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1801 |
(%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace. |
|
1802 |
(_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) |
|
1803 |
((ZCONSTR::nat |
|
1804 |
=> 'A::type |
|
1805 |
=> (nat => nat => 'A::type => bool) |
|
1806 |
=> nat => 'A::type => bool) |
|
1807 |
u ua |
|
1808 |
(%n::nat. |
|
1809 |
(_dest_rec::'A::type recspace => nat => 'A::type => bool) |
|
1810 |
(ub n))))" |
|
1811 |
||
1812 |
lemma DEF_CONSTR: "(op =::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1813 |
=> (nat |
|
1814 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1815 |
=> bool) |
|
1816 |
(CONSTR::nat |
|
1817 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1818 |
(%(u::nat) (ua::'A::type) ub::nat => 'A::type recspace. |
|
1819 |
(_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) |
|
1820 |
((ZCONSTR::nat |
|
1821 |
=> 'A::type |
|
1822 |
=> (nat => nat => 'A::type => bool) |
|
1823 |
=> nat => 'A::type => bool) |
|
1824 |
u ua |
|
1825 |
(%n::nat. |
|
1826 |
(_dest_rec::'A::type recspace => nat => 'A::type => bool) |
|
1827 |
(ub n))))" |
|
| 17644 | 1828 |
by (import hollight DEF_CONSTR) |
1829 |
||
| 17652 | 1830 |
lemma MK_REC_INJ: "(All::((nat => 'A::type => bool) => bool) => bool) |
1831 |
(%x::nat => 'A::type => bool. |
|
1832 |
(All::((nat => 'A::type => bool) => bool) => bool) |
|
1833 |
(%y::nat => 'A::type => bool. |
|
1834 |
(op -->::bool => bool => bool) |
|
1835 |
((op =::'A::type recspace => 'A::type recspace => bool) |
|
1836 |
((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) x) |
|
1837 |
((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) y)) |
|
1838 |
((op -->::bool => bool => bool) |
|
1839 |
((op &::bool => bool => bool) |
|
1840 |
((ZRECSPACE::(nat => 'A::type => bool) => bool) x) |
|
1841 |
((ZRECSPACE::(nat => 'A::type => bool) => bool) y)) |
|
1842 |
((op =::(nat => 'A::type => bool) |
|
1843 |
=> (nat => 'A::type => bool) => bool) |
|
1844 |
x y))))" |
|
| 17644 | 1845 |
by (import hollight MK_REC_INJ) |
1846 |
||
1847 |
lemma CONSTR_BOT: "ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. |
|
1848 |
CONSTR c i r ~= BOTTOM" |
|
1849 |
by (import hollight CONSTR_BOT) |
|
1850 |
||
1851 |
lemma CONSTR_INJ: "ALL (c1::nat) (i1::'A::type) (r1::nat => 'A::type recspace) (c2::nat) |
|
1852 |
(i2::'A::type) r2::nat => 'A::type recspace. |
|
1853 |
(CONSTR c1 i1 r1 = CONSTR c2 i2 r2) = (c1 = c2 & i1 = i2 & r1 = r2)" |
|
1854 |
by (import hollight CONSTR_INJ) |
|
1855 |
||
1856 |
lemma CONSTR_IND: "ALL P::'A::type recspace => bool. |
|
1857 |
P BOTTOM & |
|
1858 |
(ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. |
|
1859 |
(ALL n::nat. P (r n)) --> P (CONSTR c i r)) --> |
|
1860 |
All P" |
|
1861 |
by (import hollight CONSTR_IND) |
|
1862 |
||
1863 |
lemma CONSTR_REC: "ALL Fn::nat |
|
1864 |
=> 'A::type |
|
1865 |
=> (nat => 'A::type recspace) => (nat => 'B::type) => 'B::type. |
|
1866 |
EX f::'A::type recspace => 'B::type. |
|
1867 |
ALL (c::nat) (i::'A::type) r::nat => 'A::type recspace. |
|
1868 |
f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))" |
|
1869 |
by (import hollight CONSTR_REC) |
|
1870 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1871 |
definition FCONS :: "'A => (nat => 'A) => nat => 'A" where |
| 17644 | 1872 |
"FCONS == |
1873 |
SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type. |
|
| 17652 | 1874 |
(ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) & |
| 17644 | 1875 |
(ALL (a::'A::type) (f::nat => 'A::type) n::nat. FCONS a f (Suc n) = f n)" |
1876 |
||
1877 |
lemma DEF_FCONS: "FCONS = |
|
1878 |
(SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type. |
|
| 17652 | 1879 |
(ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) & |
| 17644 | 1880 |
(ALL (a::'A::type) (f::nat => 'A::type) n::nat. |
1881 |
FCONS a f (Suc n) = f n))" |
|
1882 |
by (import hollight DEF_FCONS) |
|
1883 |
||
| 17652 | 1884 |
lemma FCONS_UNDO: "ALL f::nat => 'A::type. f = FCONS (f 0) (f o Suc)" |
| 17644 | 1885 |
by (import hollight FCONS_UNDO) |
1886 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1887 |
definition FNIL :: "nat => 'A" where |
| 17644 | 1888 |
"FNIL == %u::nat. SOME x::'A::type. True" |
1889 |
||
1890 |
lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)" |
|
1891 |
by (import hollight DEF_FNIL) |
|
1892 |
||
1893 |
consts |
|
| 19093 | 1894 |
OUTL :: "'A + 'B => 'A" |
| 17644 | 1895 |
|
1896 |
defs |
|
1897 |
OUTL_def: "hollight.OUTL == |
|
| 19093 | 1898 |
SOME OUTL::'A::type + 'B::type => 'A::type. |
1899 |
ALL x::'A::type. OUTL (Inl x) = x" |
|
| 17644 | 1900 |
|
1901 |
lemma DEF_OUTL: "hollight.OUTL = |
|
| 19093 | 1902 |
(SOME OUTL::'A::type + 'B::type => 'A::type. |
1903 |
ALL x::'A::type. OUTL (Inl x) = x)" |
|
| 17644 | 1904 |
by (import hollight DEF_OUTL) |
1905 |
||
1906 |
consts |
|
| 19093 | 1907 |
OUTR :: "'A + 'B => 'B" |
| 17644 | 1908 |
|
1909 |
defs |
|
1910 |
OUTR_def: "hollight.OUTR == |
|
| 19093 | 1911 |
SOME OUTR::'A::type + 'B::type => 'B::type. |
1912 |
ALL y::'B::type. OUTR (Inr y) = y" |
|
| 17644 | 1913 |
|
1914 |
lemma DEF_OUTR: "hollight.OUTR = |
|
| 19093 | 1915 |
(SOME OUTR::'A::type + 'B::type => 'B::type. |
1916 |
ALL y::'B::type. OUTR (Inr y) = y)" |
|
| 17644 | 1917 |
by (import hollight DEF_OUTR) |
1918 |
||
1919 |
typedef (open) ('A) option = "(Collect::('A::type recspace => bool) => 'A::type recspace set)
|
|
1920 |
(%a::'A::type recspace. |
|
1921 |
(All::(('A::type recspace => bool) => bool) => bool)
|
|
1922 |
(%option'::'A::type recspace => bool. |
|
1923 |
(op -->::bool => bool => bool) |
|
1924 |
((All::('A::type recspace => bool) => bool)
|
|
1925 |
(%a::'A::type recspace. |
|
1926 |
(op -->::bool => bool => bool) |
|
1927 |
((op |::bool => bool => bool) |
|
1928 |
((op =::'A::type recspace => 'A::type recspace => bool) |
|
1929 |
a ((CONSTR::nat |
|
1930 |
=> 'A::type |
|
1931 |
=> (nat => 'A::type recspace) |
|
1932 |
=> 'A::type recspace) |
|
1933 |
((NUMERAL::nat => nat) (0::nat)) |
|
1934 |
((Eps::('A::type => bool) => 'A::type)
|
|
1935 |
(%v::'A::type. True::bool)) |
|
1936 |
(%n::nat. BOTTOM::'A::type recspace))) |
|
1937 |
((Ex::('A::type => bool) => bool)
|
|
1938 |
(%aa::'A::type. |
|
1939 |
(op =::'A::type recspace |
|
1940 |
=> 'A::type recspace => bool) |
|
1941 |
a ((CONSTR::nat |
|
1942 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1943 |
((Suc::nat => nat) |
|
1944 |
((NUMERAL::nat => nat) (0::nat))) |
|
1945 |
aa (%n::nat. BOTTOM::'A::type recspace))))) |
|
1946 |
(option' a))) |
|
1947 |
(option' a)))" morphisms "_dest_option" "_mk_option" |
|
1948 |
apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1949 |
((NUMERAL::nat => nat) (0::nat)) |
|
1950 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
1951 |
(%n::nat. BOTTOM::'A::type recspace)"]) |
|
1952 |
by (import hollight TYDEF_option) |
|
1953 |
||
1954 |
syntax |
|
1955 |
"_dest_option" :: _ ("'_dest'_option")
|
|
1956 |
||
1957 |
syntax |
|
1958 |
"_mk_option" :: _ ("'_mk'_option")
|
|
1959 |
||
1960 |
lemmas "TYDEF_option_@intern" = typedef_hol2hollight |
|
| 17652 | 1961 |
[where a="a :: 'A hollight.option" and r=r , |
| 17644 | 1962 |
OF type_definition_option] |
1963 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
1964 |
definition NONE :: "'A hollight.option" where |
| 17652 | 1965 |
"(op ==::'A::type hollight.option => 'A::type hollight.option => prop) |
1966 |
(NONE::'A::type hollight.option) |
|
1967 |
((_mk_option::'A::type recspace => 'A::type hollight.option) |
|
1968 |
((CONSTR::nat |
|
1969 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1970 |
(0::nat) |
|
1971 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
1972 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
1973 |
||
1974 |
lemma DEF_NONE: "(op =::'A::type hollight.option => 'A::type hollight.option => bool) |
|
1975 |
(NONE::'A::type hollight.option) |
|
1976 |
((_mk_option::'A::type recspace => 'A::type hollight.option) |
|
1977 |
((CONSTR::nat |
|
1978 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
1979 |
(0::nat) |
|
1980 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
1981 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
| 17644 | 1982 |
by (import hollight DEF_NONE) |
1983 |
||
1984 |
consts |
|
| 17652 | 1985 |
SOME :: "'A => 'A hollight.option" ("SOME")
|
| 17644 | 1986 |
|
1987 |
defs |
|
| 17652 | 1988 |
SOME_def: "(op ==::('A::type => 'A::type hollight.option)
|
1989 |
=> ('A::type => 'A::type hollight.option) => prop)
|
|
1990 |
(SOME::'A::type => 'A::type hollight.option) |
|
1991 |
(%a::'A::type. |
|
1992 |
(_mk_option::'A::type recspace => 'A::type hollight.option) |
|
1993 |
((CONSTR::nat |
|
1994 |
=> 'A::type |
|
1995 |
=> (nat => 'A::type recspace) => 'A::type recspace) |
|
1996 |
((Suc::nat => nat) (0::nat)) a |
|
1997 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
1998 |
||
1999 |
lemma DEF_SOME: "(op =::('A::type => 'A::type hollight.option)
|
|
2000 |
=> ('A::type => 'A::type hollight.option) => bool)
|
|
2001 |
(SOME::'A::type => 'A::type hollight.option) |
|
2002 |
(%a::'A::type. |
|
2003 |
(_mk_option::'A::type recspace => 'A::type hollight.option) |
|
2004 |
((CONSTR::nat |
|
2005 |
=> 'A::type |
|
2006 |
=> (nat => 'A::type recspace) => 'A::type recspace) |
|
2007 |
((Suc::nat => nat) (0::nat)) a |
|
2008 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
| 17644 | 2009 |
by (import hollight DEF_SOME) |
2010 |
||
2011 |
typedef (open) ('A) list = "(Collect::('A::type recspace => bool) => 'A::type recspace set)
|
|
2012 |
(%a::'A::type recspace. |
|
2013 |
(All::(('A::type recspace => bool) => bool) => bool)
|
|
2014 |
(%list'::'A::type recspace => bool. |
|
2015 |
(op -->::bool => bool => bool) |
|
2016 |
((All::('A::type recspace => bool) => bool)
|
|
2017 |
(%a::'A::type recspace. |
|
2018 |
(op -->::bool => bool => bool) |
|
2019 |
((op |::bool => bool => bool) |
|
2020 |
((op =::'A::type recspace => 'A::type recspace => bool) |
|
2021 |
a ((CONSTR::nat |
|
2022 |
=> 'A::type |
|
2023 |
=> (nat => 'A::type recspace) |
|
2024 |
=> 'A::type recspace) |
|
2025 |
((NUMERAL::nat => nat) (0::nat)) |
|
2026 |
((Eps::('A::type => bool) => 'A::type)
|
|
2027 |
(%v::'A::type. True::bool)) |
|
2028 |
(%n::nat. BOTTOM::'A::type recspace))) |
|
2029 |
((Ex::('A::type => bool) => bool)
|
|
2030 |
(%a0::'A::type. |
|
2031 |
(Ex::('A::type recspace => bool) => bool)
|
|
2032 |
(%a1::'A::type recspace. |
|
2033 |
(op &::bool => bool => bool) |
|
2034 |
((op =::'A::type recspace |
|
2035 |
=> 'A::type recspace => bool) |
|
2036 |
a ((CONSTR::nat |
|
2037 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
2038 |
((Suc::nat => nat) ((NUMERAL::nat => nat) (0::nat))) a0 |
|
2039 |
((FCONS::'A::type recspace |
|
2040 |
=> (nat => 'A::type recspace) => nat => 'A::type recspace) |
|
2041 |
a1 (%n::nat. BOTTOM::'A::type recspace)))) |
|
2042 |
(list' a1))))) |
|
2043 |
(list' a))) |
|
2044 |
(list' a)))" morphisms "_dest_list" "_mk_list" |
|
2045 |
apply (rule light_ex_imp_nonempty[where t="(CONSTR::nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
2046 |
((NUMERAL::nat => nat) (0::nat)) |
|
2047 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
2048 |
(%n::nat. BOTTOM::'A::type recspace)"]) |
|
2049 |
by (import hollight TYDEF_list) |
|
2050 |
||
2051 |
syntax |
|
2052 |
"_dest_list" :: _ ("'_dest'_list")
|
|
2053 |
||
2054 |
syntax |
|
2055 |
"_mk_list" :: _ ("'_mk'_list")
|
|
2056 |
||
2057 |
lemmas "TYDEF_list_@intern" = typedef_hol2hollight |
|
| 17652 | 2058 |
[where a="a :: 'A hollight.list" and r=r , |
| 17644 | 2059 |
OF type_definition_list] |
2060 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2061 |
definition NIL :: "'A hollight.list" where |
| 17652 | 2062 |
"(op ==::'A::type hollight.list => 'A::type hollight.list => prop) |
2063 |
(NIL::'A::type hollight.list) |
|
2064 |
((_mk_list::'A::type recspace => 'A::type hollight.list) |
|
2065 |
((CONSTR::nat |
|
2066 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
2067 |
(0::nat) |
|
2068 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
2069 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
2070 |
||
2071 |
lemma DEF_NIL: "(op =::'A::type hollight.list => 'A::type hollight.list => bool) |
|
2072 |
(NIL::'A::type hollight.list) |
|
2073 |
((_mk_list::'A::type recspace => 'A::type hollight.list) |
|
2074 |
((CONSTR::nat |
|
2075 |
=> 'A::type => (nat => 'A::type recspace) => 'A::type recspace) |
|
2076 |
(0::nat) |
|
2077 |
((Eps::('A::type => bool) => 'A::type) (%v::'A::type. True::bool))
|
|
2078 |
(%n::nat. BOTTOM::'A::type recspace)))" |
|
| 17644 | 2079 |
by (import hollight DEF_NIL) |
2080 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2081 |
definition CONS :: "'A => 'A hollight.list => 'A hollight.list" where |
| 17652 | 2082 |
"(op ==::('A::type => 'A::type hollight.list => 'A::type hollight.list)
|
2083 |
=> ('A::type => 'A::type hollight.list => 'A::type hollight.list)
|
|
2084 |
=> prop) |
|
2085 |
(CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list) |
|
2086 |
(%(a0::'A::type) a1::'A::type hollight.list. |
|
2087 |
(_mk_list::'A::type recspace => 'A::type hollight.list) |
|
2088 |
((CONSTR::nat |
|
2089 |
=> 'A::type |
|
2090 |
=> (nat => 'A::type recspace) => 'A::type recspace) |
|
2091 |
((Suc::nat => nat) (0::nat)) a0 |
|
2092 |
((FCONS::'A::type recspace |
|
2093 |
=> (nat => 'A::type recspace) => nat => 'A::type recspace) |
|
2094 |
((_dest_list::'A::type hollight.list => 'A::type recspace) a1) |
|
2095 |
(%n::nat. BOTTOM::'A::type recspace))))" |
|
2096 |
||
2097 |
lemma DEF_CONS: "(op =::('A::type => 'A::type hollight.list => 'A::type hollight.list)
|
|
2098 |
=> ('A::type => 'A::type hollight.list => 'A::type hollight.list)
|
|
2099 |
=> bool) |
|
2100 |
(CONS::'A::type => 'A::type hollight.list => 'A::type hollight.list) |
|
2101 |
(%(a0::'A::type) a1::'A::type hollight.list. |
|
2102 |
(_mk_list::'A::type recspace => 'A::type hollight.list) |
|
2103 |
((CONSTR::nat |
|
2104 |
=> 'A::type |
|
2105 |
=> (nat => 'A::type recspace) => 'A::type recspace) |
|
2106 |
((Suc::nat => nat) (0::nat)) a0 |
|
2107 |
((FCONS::'A::type recspace |
|
2108 |
=> (nat => 'A::type recspace) => nat => 'A::type recspace) |
|
2109 |
((_dest_list::'A::type hollight.list => 'A::type recspace) a1) |
|
2110 |
(%n::nat. BOTTOM::'A::type recspace))))" |
|
| 17644 | 2111 |
by (import hollight DEF_CONS) |
2112 |
||
2113 |
lemma pair_RECURSION: "ALL PAIR'::'A::type => 'B::type => 'C::type. |
|
2114 |
EX x::'A::type * 'B::type => 'C::type. |
|
2115 |
ALL (a0::'A::type) a1::'B::type. x (a0, a1) = PAIR' a0 a1" |
|
2116 |
by (import hollight pair_RECURSION) |
|
2117 |
||
2118 |
lemma num_RECURSION_STD: "ALL (e::'Z::type) f::nat => 'Z::type => 'Z::type. |
|
| 17652 | 2119 |
EX fn::nat => 'Z::type. fn 0 = e & (ALL n::nat. fn (Suc n) = f n (fn n))" |
| 17644 | 2120 |
by (import hollight num_RECURSION_STD) |
2121 |
||
| 19093 | 2122 |
lemma bool_RECURSION: "ALL (a::'A::type) b::'A::type. |
2123 |
EX x::bool => 'A::type. x False = a & x True = b" |
|
2124 |
by (import hollight bool_RECURSION) |
|
2125 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2126 |
definition ISO :: "('A => 'B) => ('B => 'A) => bool" where
|
| 17644 | 2127 |
"ISO == |
2128 |
%(u::'A::type => 'B::type) ua::'B::type => 'A::type. |
|
2129 |
(ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y)" |
|
2130 |
||
2131 |
lemma DEF_ISO: "ISO = |
|
2132 |
(%(u::'A::type => 'B::type) ua::'B::type => 'A::type. |
|
2133 |
(ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y))" |
|
2134 |
by (import hollight DEF_ISO) |
|
2135 |
||
2136 |
lemma ISO_REFL: "ISO (%x::'A::type. x) (%x::'A::type. x)" |
|
2137 |
by (import hollight ISO_REFL) |
|
2138 |
||
2139 |
lemma ISO_FUN: "ISO (f::'A::type => 'A'::type) (f'::'A'::type => 'A::type) & |
|
2140 |
ISO (g::'B::type => 'B'::type) (g'::'B'::type => 'B::type) --> |
|
2141 |
ISO (%(h::'A::type => 'B::type) a'::'A'::type. g (h (f' a'))) |
|
2142 |
(%(h::'A'::type => 'B'::type) a::'A::type. g' (h (f a)))" |
|
2143 |
by (import hollight ISO_FUN) |
|
2144 |
||
| 19093 | 2145 |
lemma ISO_USAGE: "ISO (f::'q_16636::type => 'q_16633::type) |
2146 |
(g::'q_16633::type => 'q_16636::type) --> |
|
2147 |
(ALL P::'q_16636::type => bool. All P = (ALL x::'q_16633::type. P (g x))) & |
|
2148 |
(ALL P::'q_16636::type => bool. Ex P = (EX x::'q_16633::type. P (g x))) & |
|
2149 |
(ALL (a::'q_16636::type) b::'q_16633::type. (a = g b) = (f a = b))" |
|
| 17644 | 2150 |
by (import hollight ISO_USAGE) |
2151 |
||
2152 |
typedef (open) N_2 = "{a::bool recspace.
|
|
2153 |
ALL u::bool recspace => bool. |
|
2154 |
(ALL a::bool recspace. |
|
| 17652 | 2155 |
a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) | |
| 17644 | 2156 |
a = |
| 17652 | 2157 |
CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) --> |
| 17644 | 2158 |
u a) --> |
2159 |
u a}" morphisms "_dest_2" "_mk_2" |
|
| 17652 | 2160 |
apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"]) |
| 17644 | 2161 |
by (import hollight TYDEF_2) |
2162 |
||
2163 |
syntax |
|
2164 |
"_dest_2" :: _ ("'_dest'_2")
|
|
2165 |
||
2166 |
syntax |
|
2167 |
"_mk_2" :: _ ("'_mk'_2")
|
|
2168 |
||
2169 |
lemmas "TYDEF_2_@intern" = typedef_hol2hollight |
|
2170 |
[where a="a :: N_2" and r=r , |
|
2171 |
OF type_definition_N_2] |
|
2172 |
||
2173 |
consts |
|
| 19093 | 2174 |
"_10302" :: "N_2" ("'_10302")
|
| 17644 | 2175 |
|
2176 |
defs |
|
| 19093 | 2177 |
"_10302_def": "(op ==::N_2 => N_2 => prop) (_10302::N_2) |
| 17652 | 2178 |
((_mk_2::bool recspace => N_2) |
2179 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2180 |
(0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2181 |
(%n::nat. BOTTOM::bool recspace)))" |
|
2182 |
||
| 19093 | 2183 |
lemma DEF__10302: "(op =::N_2 => N_2 => bool) (_10302::N_2) |
| 17652 | 2184 |
((_mk_2::bool recspace => N_2) |
2185 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2186 |
(0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2187 |
(%n::nat. BOTTOM::bool recspace)))" |
|
| 19093 | 2188 |
by (import hollight DEF__10302) |
| 17644 | 2189 |
|
2190 |
consts |
|
| 19093 | 2191 |
"_10303" :: "N_2" ("'_10303")
|
| 17644 | 2192 |
|
2193 |
defs |
|
| 19093 | 2194 |
"_10303_def": "(op ==::N_2 => N_2 => prop) (_10303::N_2) |
| 17652 | 2195 |
((_mk_2::bool recspace => N_2) |
2196 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2197 |
((Suc::nat => nat) (0::nat)) |
|
2198 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2199 |
(%n::nat. BOTTOM::bool recspace)))" |
|
2200 |
||
| 19093 | 2201 |
lemma DEF__10303: "(op =::N_2 => N_2 => bool) (_10303::N_2) |
| 17652 | 2202 |
((_mk_2::bool recspace => N_2) |
2203 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2204 |
((Suc::nat => nat) (0::nat)) |
|
2205 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2206 |
(%n::nat. BOTTOM::bool recspace)))" |
|
| 19093 | 2207 |
by (import hollight DEF__10303) |
| 17644 | 2208 |
|
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2209 |
definition two_1 :: "N_2" where |
| 19093 | 2210 |
"two_1 == _10302" |
2211 |
||
2212 |
lemma DEF_two_1: "two_1 = _10302" |
|
| 17644 | 2213 |
by (import hollight DEF_two_1) |
2214 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2215 |
definition two_2 :: "N_2" where |
| 19093 | 2216 |
"two_2 == _10303" |
2217 |
||
2218 |
lemma DEF_two_2: "two_2 = _10303" |
|
| 17644 | 2219 |
by (import hollight DEF_two_2) |
2220 |
||
2221 |
typedef (open) N_3 = "{a::bool recspace.
|
|
2222 |
ALL u::bool recspace => bool. |
|
2223 |
(ALL a::bool recspace. |
|
| 17652 | 2224 |
a = CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM) | |
| 17644 | 2225 |
a = |
| 17652 | 2226 |
CONSTR (Suc (NUMERAL 0)) (SOME x::bool. True) (%n::nat. BOTTOM) | |
| 17644 | 2227 |
a = |
| 17652 | 2228 |
CONSTR (Suc (Suc (NUMERAL 0))) (SOME x::bool. True) |
| 17644 | 2229 |
(%n::nat. BOTTOM) --> |
2230 |
u a) --> |
|
2231 |
u a}" morphisms "_dest_3" "_mk_3" |
|
| 17652 | 2232 |
apply (rule light_ex_imp_nonempty[where t="CONSTR (NUMERAL 0) (SOME x::bool. True) (%n::nat. BOTTOM)"]) |
| 17644 | 2233 |
by (import hollight TYDEF_3) |
2234 |
||
2235 |
syntax |
|
2236 |
"_dest_3" :: _ ("'_dest'_3")
|
|
2237 |
||
2238 |
syntax |
|
2239 |
"_mk_3" :: _ ("'_mk'_3")
|
|
2240 |
||
2241 |
lemmas "TYDEF_3_@intern" = typedef_hol2hollight |
|
2242 |
[where a="a :: N_3" and r=r , |
|
2243 |
OF type_definition_N_3] |
|
2244 |
||
2245 |
consts |
|
| 19093 | 2246 |
"_10326" :: "N_3" ("'_10326")
|
| 17644 | 2247 |
|
2248 |
defs |
|
| 19093 | 2249 |
"_10326_def": "(op ==::N_3 => N_3 => prop) (_10326::N_3) |
| 17652 | 2250 |
((_mk_3::bool recspace => N_3) |
2251 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2252 |
(0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2253 |
(%n::nat. BOTTOM::bool recspace)))" |
|
2254 |
||
| 19093 | 2255 |
lemma DEF__10326: "(op =::N_3 => N_3 => bool) (_10326::N_3) |
| 17652 | 2256 |
((_mk_3::bool recspace => N_3) |
2257 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2258 |
(0::nat) ((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2259 |
(%n::nat. BOTTOM::bool recspace)))" |
|
| 19093 | 2260 |
by (import hollight DEF__10326) |
| 17644 | 2261 |
|
2262 |
consts |
|
| 19093 | 2263 |
"_10327" :: "N_3" ("'_10327")
|
| 17644 | 2264 |
|
2265 |
defs |
|
| 19093 | 2266 |
"_10327_def": "(op ==::N_3 => N_3 => prop) (_10327::N_3) |
| 17652 | 2267 |
((_mk_3::bool recspace => N_3) |
2268 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2269 |
((Suc::nat => nat) (0::nat)) |
|
2270 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2271 |
(%n::nat. BOTTOM::bool recspace)))" |
|
2272 |
||
| 19093 | 2273 |
lemma DEF__10327: "(op =::N_3 => N_3 => bool) (_10327::N_3) |
| 17652 | 2274 |
((_mk_3::bool recspace => N_3) |
2275 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2276 |
((Suc::nat => nat) (0::nat)) |
|
2277 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2278 |
(%n::nat. BOTTOM::bool recspace)))" |
|
| 19093 | 2279 |
by (import hollight DEF__10327) |
| 17644 | 2280 |
|
2281 |
consts |
|
| 19093 | 2282 |
"_10328" :: "N_3" ("'_10328")
|
| 17644 | 2283 |
|
2284 |
defs |
|
| 19093 | 2285 |
"_10328_def": "(op ==::N_3 => N_3 => prop) (_10328::N_3) |
| 17652 | 2286 |
((_mk_3::bool recspace => N_3) |
2287 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2288 |
((Suc::nat => nat) ((Suc::nat => nat) (0::nat))) |
|
2289 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2290 |
(%n::nat. BOTTOM::bool recspace)))" |
|
2291 |
||
| 19093 | 2292 |
lemma DEF__10328: "(op =::N_3 => N_3 => bool) (_10328::N_3) |
| 17652 | 2293 |
((_mk_3::bool recspace => N_3) |
2294 |
((CONSTR::nat => bool => (nat => bool recspace) => bool recspace) |
|
2295 |
((Suc::nat => nat) ((Suc::nat => nat) (0::nat))) |
|
2296 |
((Eps::(bool => bool) => bool) (%x::bool. True::bool)) |
|
2297 |
(%n::nat. BOTTOM::bool recspace)))" |
|
| 19093 | 2298 |
by (import hollight DEF__10328) |
| 17644 | 2299 |
|
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2300 |
definition three_1 :: "N_3" where |
| 19093 | 2301 |
"three_1 == _10326" |
2302 |
||
2303 |
lemma DEF_three_1: "three_1 = _10326" |
|
| 17644 | 2304 |
by (import hollight DEF_three_1) |
2305 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2306 |
definition three_2 :: "N_3" where |
| 19093 | 2307 |
"three_2 == _10327" |
2308 |
||
2309 |
lemma DEF_three_2: "three_2 = _10327" |
|
| 17644 | 2310 |
by (import hollight DEF_three_2) |
2311 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2312 |
definition three_3 :: "N_3" where |
| 19093 | 2313 |
"three_3 == _10328" |
2314 |
||
2315 |
lemma DEF_three_3: "three_3 = _10328" |
|
| 17644 | 2316 |
by (import hollight DEF_three_3) |
2317 |
||
2318 |
lemma list_INDUCT: "ALL P::'A::type hollight.list => bool. |
|
2319 |
P NIL & |
|
2320 |
(ALL (a0::'A::type) a1::'A::type hollight.list. |
|
2321 |
P a1 --> P (CONS a0 a1)) --> |
|
2322 |
All P" |
|
2323 |
by (import hollight list_INDUCT) |
|
2324 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2325 |
definition HD :: "'A hollight.list => 'A" where |
| 17644 | 2326 |
"HD == |
2327 |
SOME HD::'A::type hollight.list => 'A::type. |
|
2328 |
ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h" |
|
2329 |
||
2330 |
lemma DEF_HD: "HD = |
|
2331 |
(SOME HD::'A::type hollight.list => 'A::type. |
|
2332 |
ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h)" |
|
2333 |
by (import hollight DEF_HD) |
|
2334 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2335 |
definition TL :: "'A hollight.list => 'A hollight.list" where |
| 17644 | 2336 |
"TL == |
2337 |
SOME TL::'A::type hollight.list => 'A::type hollight.list. |
|
2338 |
ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t" |
|
2339 |
||
2340 |
lemma DEF_TL: "TL = |
|
2341 |
(SOME TL::'A::type hollight.list => 'A::type hollight.list. |
|
2342 |
ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t)" |
|
2343 |
by (import hollight DEF_TL) |
|
2344 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2345 |
definition APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" where |
| 17644 | 2346 |
"APPEND == |
2347 |
SOME APPEND::'A::type hollight.list |
|
2348 |
=> 'A::type hollight.list => 'A::type hollight.list. |
|
2349 |
(ALL l::'A::type hollight.list. APPEND NIL l = l) & |
|
2350 |
(ALL (h::'A::type) (t::'A::type hollight.list) l::'A::type hollight.list. |
|
2351 |
APPEND (CONS h t) l = CONS h (APPEND t l))" |
|
2352 |
||
2353 |
lemma DEF_APPEND: "APPEND = |
|
2354 |
(SOME APPEND::'A::type hollight.list |
|
2355 |
=> 'A::type hollight.list => 'A::type hollight.list. |
|
2356 |
(ALL l::'A::type hollight.list. APPEND NIL l = l) & |
|
2357 |
(ALL (h::'A::type) (t::'A::type hollight.list) |
|
2358 |
l::'A::type hollight.list. |
|
2359 |
APPEND (CONS h t) l = CONS h (APPEND t l)))" |
|
2360 |
by (import hollight DEF_APPEND) |
|
2361 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2362 |
definition REVERSE :: "'A hollight.list => 'A hollight.list" where |
| 17644 | 2363 |
"REVERSE == |
2364 |
SOME REVERSE::'A::type hollight.list => 'A::type hollight.list. |
|
2365 |
REVERSE NIL = NIL & |
|
2366 |
(ALL (l::'A::type hollight.list) x::'A::type. |
|
2367 |
REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL))" |
|
2368 |
||
2369 |
lemma DEF_REVERSE: "REVERSE = |
|
2370 |
(SOME REVERSE::'A::type hollight.list => 'A::type hollight.list. |
|
2371 |
REVERSE NIL = NIL & |
|
2372 |
(ALL (l::'A::type hollight.list) x::'A::type. |
|
2373 |
REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL)))" |
|
2374 |
by (import hollight DEF_REVERSE) |
|
2375 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2376 |
definition LENGTH :: "'A hollight.list => nat" where |
| 17644 | 2377 |
"LENGTH == |
2378 |
SOME LENGTH::'A::type hollight.list => nat. |
|
| 17652 | 2379 |
LENGTH NIL = 0 & |
| 17644 | 2380 |
(ALL (h::'A::type) t::'A::type hollight.list. |
2381 |
LENGTH (CONS h t) = Suc (LENGTH t))" |
|
2382 |
||
2383 |
lemma DEF_LENGTH: "LENGTH = |
|
2384 |
(SOME LENGTH::'A::type hollight.list => nat. |
|
| 17652 | 2385 |
LENGTH NIL = 0 & |
| 17644 | 2386 |
(ALL (h::'A::type) t::'A::type hollight.list. |
2387 |
LENGTH (CONS h t) = Suc (LENGTH t)))" |
|
2388 |
by (import hollight DEF_LENGTH) |
|
2389 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2390 |
definition MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" where
|
| 17644 | 2391 |
"MAP == |
2392 |
SOME MAP::('A::type => 'B::type)
|
|
2393 |
=> 'A::type hollight.list => 'B::type hollight.list. |
|
2394 |
(ALL f::'A::type => 'B::type. MAP f NIL = NIL) & |
|
2395 |
(ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list. |
|
2396 |
MAP f (CONS h t) = CONS (f h) (MAP f t))" |
|
2397 |
||
2398 |
lemma DEF_MAP: "MAP = |
|
2399 |
(SOME MAP::('A::type => 'B::type)
|
|
2400 |
=> 'A::type hollight.list => 'B::type hollight.list. |
|
2401 |
(ALL f::'A::type => 'B::type. MAP f NIL = NIL) & |
|
2402 |
(ALL (f::'A::type => 'B::type) (h::'A::type) t::'A::type hollight.list. |
|
2403 |
MAP f (CONS h t) = CONS (f h) (MAP f t)))" |
|
2404 |
by (import hollight DEF_MAP) |
|
2405 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2406 |
definition LAST :: "'A hollight.list => 'A" where |
| 17644 | 2407 |
"LAST == |
2408 |
SOME LAST::'A::type hollight.list => 'A::type. |
|
2409 |
ALL (h::'A::type) t::'A::type hollight.list. |
|
2410 |
LAST (CONS h t) = COND (t = NIL) h (LAST t)" |
|
2411 |
||
2412 |
lemma DEF_LAST: "LAST = |
|
2413 |
(SOME LAST::'A::type hollight.list => 'A::type. |
|
2414 |
ALL (h::'A::type) t::'A::type hollight.list. |
|
2415 |
LAST (CONS h t) = COND (t = NIL) h (LAST t))" |
|
2416 |
by (import hollight DEF_LAST) |
|
2417 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2418 |
definition REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" where |
| 17644 | 2419 |
"REPLICATE == |
| 19093 | 2420 |
SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list. |
2421 |
(ALL x::'q_16860::type. REPLICATE 0 x = NIL) & |
|
2422 |
(ALL (n::nat) x::'q_16860::type. |
|
| 17644 | 2423 |
REPLICATE (Suc n) x = CONS x (REPLICATE n x))" |
2424 |
||
2425 |
lemma DEF_REPLICATE: "REPLICATE = |
|
| 19093 | 2426 |
(SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list. |
2427 |
(ALL x::'q_16860::type. REPLICATE 0 x = NIL) & |
|
2428 |
(ALL (n::nat) x::'q_16860::type. |
|
| 17644 | 2429 |
REPLICATE (Suc n) x = CONS x (REPLICATE n x)))" |
2430 |
by (import hollight DEF_REPLICATE) |
|
2431 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2432 |
definition NULL :: "'q_16875 hollight.list => bool" where |
| 17644 | 2433 |
"NULL == |
| 19093 | 2434 |
SOME NULL::'q_16875::type hollight.list => bool. |
| 17644 | 2435 |
NULL NIL = True & |
| 19093 | 2436 |
(ALL (h::'q_16875::type) t::'q_16875::type hollight.list. |
| 17644 | 2437 |
NULL (CONS h t) = False)" |
2438 |
||
2439 |
lemma DEF_NULL: "NULL = |
|
| 19093 | 2440 |
(SOME NULL::'q_16875::type hollight.list => bool. |
| 17644 | 2441 |
NULL NIL = True & |
| 19093 | 2442 |
(ALL (h::'q_16875::type) t::'q_16875::type hollight.list. |
| 17644 | 2443 |
NULL (CONS h t) = False))" |
2444 |
by (import hollight DEF_NULL) |
|
2445 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2446 |
definition ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" where
|
| 17644 | 2447 |
"ALL_list == |
| 19093 | 2448 |
SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool.
|
2449 |
(ALL P::'q_16895::type => bool. u P NIL = True) & |
|
2450 |
(ALL (h::'q_16895::type) (P::'q_16895::type => bool) |
|
2451 |
t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t))" |
|
| 17644 | 2452 |
|
2453 |
lemma DEF_ALL: "ALL_list = |
|
| 19093 | 2454 |
(SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool.
|
2455 |
(ALL P::'q_16895::type => bool. u P NIL = True) & |
|
2456 |
(ALL (h::'q_16895::type) (P::'q_16895::type => bool) |
|
2457 |
t::'q_16895::type hollight.list. u P (CONS h t) = (P h & u P t)))" |
|
| 17644 | 2458 |
by (import hollight DEF_ALL) |
2459 |
||
2460 |
consts |
|
| 19093 | 2461 |
EX :: "('q_16916 => bool) => 'q_16916 hollight.list => bool" ("EX")
|
| 17644 | 2462 |
|
2463 |
defs |
|
2464 |
EX_def: "EX == |
|
| 19093 | 2465 |
SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool.
|
2466 |
(ALL P::'q_16916::type => bool. u P NIL = False) & |
|
2467 |
(ALL (h::'q_16916::type) (P::'q_16916::type => bool) |
|
2468 |
t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t))" |
|
| 17644 | 2469 |
|
2470 |
lemma DEF_EX: "EX = |
|
| 19093 | 2471 |
(SOME u::('q_16916::type => bool) => 'q_16916::type hollight.list => bool.
|
2472 |
(ALL P::'q_16916::type => bool. u P NIL = False) & |
|
2473 |
(ALL (h::'q_16916::type) (P::'q_16916::type => bool) |
|
2474 |
t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t)))" |
|
| 17644 | 2475 |
by (import hollight DEF_EX) |
2476 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2477 |
definition ITLIST :: "('q_16939 => 'q_16938 => 'q_16938)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2478 |
=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" where |
| 17644 | 2479 |
"ITLIST == |
| 19093 | 2480 |
SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type)
|
2481 |
=> 'q_16939::type hollight.list |
|
2482 |
=> 'q_16938::type => 'q_16938::type. |
|
2483 |
(ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type) |
|
2484 |
b::'q_16938::type. ITLIST f NIL b = b) & |
|
2485 |
(ALL (h::'q_16939::type) |
|
2486 |
(f::'q_16939::type => 'q_16938::type => 'q_16938::type) |
|
2487 |
(t::'q_16939::type hollight.list) b::'q_16938::type. |
|
| 17644 | 2488 |
ITLIST f (CONS h t) b = f h (ITLIST f t b))" |
2489 |
||
2490 |
lemma DEF_ITLIST: "ITLIST = |
|
| 19093 | 2491 |
(SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type)
|
2492 |
=> 'q_16939::type hollight.list |
|
2493 |
=> 'q_16938::type => 'q_16938::type. |
|
2494 |
(ALL (f::'q_16939::type => 'q_16938::type => 'q_16938::type) |
|
2495 |
b::'q_16938::type. ITLIST f NIL b = b) & |
|
2496 |
(ALL (h::'q_16939::type) |
|
2497 |
(f::'q_16939::type => 'q_16938::type => 'q_16938::type) |
|
2498 |
(t::'q_16939::type hollight.list) b::'q_16938::type. |
|
| 17644 | 2499 |
ITLIST f (CONS h t) b = f h (ITLIST f t b)))" |
2500 |
by (import hollight DEF_ITLIST) |
|
2501 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2502 |
definition MEM :: "'q_16964 => 'q_16964 hollight.list => bool" where |
| 17644 | 2503 |
"MEM == |
| 19093 | 2504 |
SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool. |
2505 |
(ALL x::'q_16964::type. MEM x NIL = False) & |
|
2506 |
(ALL (h::'q_16964::type) (x::'q_16964::type) |
|
2507 |
t::'q_16964::type hollight.list. |
|
| 17644 | 2508 |
MEM x (CONS h t) = (x = h | MEM x t))" |
2509 |
||
2510 |
lemma DEF_MEM: "MEM = |
|
| 19093 | 2511 |
(SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool. |
2512 |
(ALL x::'q_16964::type. MEM x NIL = False) & |
|
2513 |
(ALL (h::'q_16964::type) (x::'q_16964::type) |
|
2514 |
t::'q_16964::type hollight.list. |
|
| 17644 | 2515 |
MEM x (CONS h t) = (x = h | MEM x t)))" |
2516 |
by (import hollight DEF_MEM) |
|
2517 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2518 |
definition ALL2 :: "('q_16997 => 'q_17004 => bool)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2519 |
=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" where |
| 17644 | 2520 |
"ALL2 == |
| 19093 | 2521 |
SOME ALL2::('q_16997::type => 'q_17004::type => bool)
|
2522 |
=> 'q_16997::type hollight.list |
|
2523 |
=> 'q_17004::type hollight.list => bool. |
|
2524 |
(ALL (P::'q_16997::type => 'q_17004::type => bool) |
|
2525 |
l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & |
|
2526 |
(ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool) |
|
2527 |
(t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list. |
|
| 17644 | 2528 |
ALL2 P (CONS h1 t1) l2 = |
2529 |
COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2)))" |
|
2530 |
||
2531 |
lemma DEF_ALL2: "ALL2 = |
|
| 19093 | 2532 |
(SOME ALL2::('q_16997::type => 'q_17004::type => bool)
|
2533 |
=> 'q_16997::type hollight.list |
|
2534 |
=> 'q_17004::type hollight.list => bool. |
|
2535 |
(ALL (P::'q_16997::type => 'q_17004::type => bool) |
|
2536 |
l2::'q_17004::type hollight.list. ALL2 P NIL l2 = (l2 = NIL)) & |
|
2537 |
(ALL (h1::'q_16997::type) (P::'q_16997::type => 'q_17004::type => bool) |
|
2538 |
(t1::'q_16997::type hollight.list) l2::'q_17004::type hollight.list. |
|
| 17644 | 2539 |
ALL2 P (CONS h1 t1) l2 = |
2540 |
COND (l2 = NIL) False (P h1 (HD l2) & ALL2 P t1 (TL l2))))" |
|
2541 |
by (import hollight DEF_ALL2) |
|
2542 |
||
| 19093 | 2543 |
lemma ALL2: "ALL2 (P::'q_17059::type => 'q_17058::type => bool) NIL NIL = True & |
2544 |
ALL2 P (CONS (h1::'q_17059::type) (t1::'q_17059::type hollight.list)) NIL = |
|
| 17644 | 2545 |
False & |
| 19093 | 2546 |
ALL2 P NIL (CONS (h2::'q_17058::type) (t2::'q_17058::type hollight.list)) = |
| 17644 | 2547 |
False & |
2548 |
ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)" |
|
2549 |
by (import hollight ALL2) |
|
2550 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2551 |
definition MAP2 :: "('q_17089 => 'q_17096 => 'q_17086)
|
| 19093 | 2552 |
=> 'q_17089 hollight.list |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2553 |
=> 'q_17096 hollight.list => 'q_17086 hollight.list" where |
| 17644 | 2554 |
"MAP2 == |
| 19093 | 2555 |
SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type)
|
2556 |
=> 'q_17089::type hollight.list |
|
2557 |
=> 'q_17096::type hollight.list |
|
2558 |
=> 'q_17086::type hollight.list. |
|
2559 |
(ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type) |
|
2560 |
l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) & |
|
2561 |
(ALL (h1::'q_17089::type) |
|
2562 |
(f::'q_17089::type => 'q_17096::type => 'q_17086::type) |
|
2563 |
(t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list. |
|
| 17644 | 2564 |
MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l)))" |
2565 |
||
2566 |
lemma DEF_MAP2: "MAP2 = |
|
| 19093 | 2567 |
(SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type)
|
2568 |
=> 'q_17089::type hollight.list |
|
2569 |
=> 'q_17096::type hollight.list |
|
2570 |
=> 'q_17086::type hollight.list. |
|
2571 |
(ALL (f::'q_17089::type => 'q_17096::type => 'q_17086::type) |
|
2572 |
l::'q_17096::type hollight.list. MAP2 f NIL l = NIL) & |
|
2573 |
(ALL (h1::'q_17089::type) |
|
2574 |
(f::'q_17089::type => 'q_17096::type => 'q_17086::type) |
|
2575 |
(t1::'q_17089::type hollight.list) l::'q_17096::type hollight.list. |
|
| 17644 | 2576 |
MAP2 f (CONS h1 t1) l = CONS (f h1 (HD l)) (MAP2 f t1 (TL l))))" |
2577 |
by (import hollight DEF_MAP2) |
|
2578 |
||
| 19093 | 2579 |
lemma MAP2: "MAP2 (f::'q_17131::type => 'q_17130::type => 'q_17137::type) NIL NIL = NIL & |
2580 |
MAP2 f (CONS (h1::'q_17131::type) (t1::'q_17131::type hollight.list)) |
|
2581 |
(CONS (h2::'q_17130::type) (t2::'q_17130::type hollight.list)) = |
|
| 17644 | 2582 |
CONS (f h1 h2) (MAP2 f t1 t2)" |
2583 |
by (import hollight MAP2) |
|
2584 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2585 |
definition EL :: "nat => 'q_17157 hollight.list => 'q_17157" where |
| 17644 | 2586 |
"EL == |
| 19093 | 2587 |
SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type. |
2588 |
(ALL l::'q_17157::type hollight.list. EL 0 l = HD l) & |
|
2589 |
(ALL (n::nat) l::'q_17157::type hollight.list. |
|
| 17644 | 2590 |
EL (Suc n) l = EL n (TL l))" |
2591 |
||
2592 |
lemma DEF_EL: "EL = |
|
| 19093 | 2593 |
(SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type. |
2594 |
(ALL l::'q_17157::type hollight.list. EL 0 l = HD l) & |
|
2595 |
(ALL (n::nat) l::'q_17157::type hollight.list. |
|
| 17644 | 2596 |
EL (Suc n) l = EL n (TL l)))" |
2597 |
by (import hollight DEF_EL) |
|
2598 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2599 |
definition FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" where
|
| 17644 | 2600 |
"FILTER == |
| 19093 | 2601 |
SOME FILTER::('q_17182::type => bool)
|
2602 |
=> 'q_17182::type hollight.list |
|
2603 |
=> 'q_17182::type hollight.list. |
|
2604 |
(ALL P::'q_17182::type => bool. FILTER P NIL = NIL) & |
|
2605 |
(ALL (h::'q_17182::type) (P::'q_17182::type => bool) |
|
2606 |
t::'q_17182::type hollight.list. |
|
| 17644 | 2607 |
FILTER P (CONS h t) = COND (P h) (CONS h (FILTER P t)) (FILTER P t))" |
2608 |
||
2609 |
lemma DEF_FILTER: "FILTER = |
|
| 19093 | 2610 |
(SOME FILTER::('q_17182::type => bool)
|
2611 |
=> 'q_17182::type hollight.list |
|
2612 |
=> 'q_17182::type hollight.list. |
|
2613 |
(ALL P::'q_17182::type => bool. FILTER P NIL = NIL) & |
|
2614 |
(ALL (h::'q_17182::type) (P::'q_17182::type => bool) |
|
2615 |
t::'q_17182::type hollight.list. |
|
| 17644 | 2616 |
FILTER P (CONS h t) = |
2617 |
COND (P h) (CONS h (FILTER P t)) (FILTER P t)))" |
|
2618 |
by (import hollight DEF_FILTER) |
|
2619 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2620 |
definition ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" where
|
| 17644 | 2621 |
"ASSOC == |
| 19093 | 2622 |
SOME ASSOC::'q_17211::type |
2623 |
=> ('q_17211::type * 'q_17205::type) hollight.list
|
|
2624 |
=> 'q_17205::type. |
|
2625 |
ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type) |
|
2626 |
t::('q_17211::type * 'q_17205::type) hollight.list.
|
|
| 17644 | 2627 |
ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t)" |
2628 |
||
2629 |
lemma DEF_ASSOC: "ASSOC = |
|
| 19093 | 2630 |
(SOME ASSOC::'q_17211::type |
2631 |
=> ('q_17211::type * 'q_17205::type) hollight.list
|
|
2632 |
=> 'q_17205::type. |
|
2633 |
ALL (h::'q_17211::type * 'q_17205::type) (a::'q_17211::type) |
|
2634 |
t::('q_17211::type * 'q_17205::type) hollight.list.
|
|
| 17644 | 2635 |
ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))" |
2636 |
by (import hollight DEF_ASSOC) |
|
2637 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2638 |
definition ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2639 |
=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" where |
| 17644 | 2640 |
"ITLIST2 == |
| 19093 | 2641 |
SOME ITLIST2::('q_17235::type
|
2642 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2643 |
=> 'q_17235::type hollight.list |
|
2644 |
=> 'q_17243::type hollight.list |
|
2645 |
=> 'q_17233::type => 'q_17233::type. |
|
2646 |
(ALL (f::'q_17235::type |
|
2647 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2648 |
(l2::'q_17243::type hollight.list) b::'q_17233::type. |
|
| 17644 | 2649 |
ITLIST2 f NIL l2 b = b) & |
| 19093 | 2650 |
(ALL (h1::'q_17235::type) |
2651 |
(f::'q_17235::type |
|
2652 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2653 |
(t1::'q_17235::type hollight.list) (l2::'q_17243::type hollight.list) |
|
2654 |
b::'q_17233::type. |
|
| 17644 | 2655 |
ITLIST2 f (CONS h1 t1) l2 b = f h1 (HD l2) (ITLIST2 f t1 (TL l2) b))" |
2656 |
||
2657 |
lemma DEF_ITLIST2: "ITLIST2 = |
|
| 19093 | 2658 |
(SOME ITLIST2::('q_17235::type
|
2659 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2660 |
=> 'q_17235::type hollight.list |
|
2661 |
=> 'q_17243::type hollight.list |
|
2662 |
=> 'q_17233::type => 'q_17233::type. |
|
2663 |
(ALL (f::'q_17235::type |
|
2664 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2665 |
(l2::'q_17243::type hollight.list) b::'q_17233::type. |
|
| 17644 | 2666 |
ITLIST2 f NIL l2 b = b) & |
| 19093 | 2667 |
(ALL (h1::'q_17235::type) |
2668 |
(f::'q_17235::type |
|
2669 |
=> 'q_17243::type => 'q_17233::type => 'q_17233::type) |
|
2670 |
(t1::'q_17235::type hollight.list) |
|
2671 |
(l2::'q_17243::type hollight.list) b::'q_17233::type. |
|
| 17644 | 2672 |
ITLIST2 f (CONS h1 t1) l2 b = |
2673 |
f h1 (HD l2) (ITLIST2 f t1 (TL l2) b)))" |
|
2674 |
by (import hollight DEF_ITLIST2) |
|
2675 |
||
2676 |
lemma ITLIST2: "ITLIST2 |
|
| 19093 | 2677 |
(f::'q_17277::type => 'q_17276::type => 'q_17275::type => 'q_17275::type) |
2678 |
NIL NIL (b::'q_17275::type) = |
|
| 17644 | 2679 |
b & |
| 19093 | 2680 |
ITLIST2 f (CONS (h1::'q_17277::type) (t1::'q_17277::type hollight.list)) |
2681 |
(CONS (h2::'q_17276::type) (t2::'q_17276::type hollight.list)) b = |
|
| 17644 | 2682 |
f h1 h2 (ITLIST2 f t1 t2 b)" |
2683 |
by (import hollight ITLIST2) |
|
2684 |
||
2685 |
consts |
|
| 19093 | 2686 |
ZIP :: "'q_17307 hollight.list |
2687 |
=> 'q_17315 hollight.list => ('q_17307 * 'q_17315) hollight.list"
|
|
| 17644 | 2688 |
|
2689 |
defs |
|
2690 |
ZIP_def: "hollight.ZIP == |
|
| 19093 | 2691 |
SOME ZIP::'q_17307::type hollight.list |
2692 |
=> 'q_17315::type hollight.list |
|
2693 |
=> ('q_17307::type * 'q_17315::type) hollight.list.
|
|
2694 |
(ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) & |
|
2695 |
(ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list) |
|
2696 |
l2::'q_17315::type hollight.list. |
|
| 17644 | 2697 |
ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2)))" |
2698 |
||
2699 |
lemma DEF_ZIP: "hollight.ZIP = |
|
| 19093 | 2700 |
(SOME ZIP::'q_17307::type hollight.list |
2701 |
=> 'q_17315::type hollight.list |
|
2702 |
=> ('q_17307::type * 'q_17315::type) hollight.list.
|
|
2703 |
(ALL l2::'q_17315::type hollight.list. ZIP NIL l2 = NIL) & |
|
2704 |
(ALL (h1::'q_17307::type) (t1::'q_17307::type hollight.list) |
|
2705 |
l2::'q_17315::type hollight.list. |
|
| 17644 | 2706 |
ZIP (CONS h1 t1) l2 = CONS (h1, HD l2) (ZIP t1 (TL l2))))" |
2707 |
by (import hollight DEF_ZIP) |
|
2708 |
||
2709 |
lemma ZIP: "(op &::bool => bool => bool) |
|
| 19093 | 2710 |
((op =::('q_17326::type * 'q_17327::type) hollight.list
|
2711 |
=> ('q_17326::type * 'q_17327::type) hollight.list => bool)
|
|
2712 |
((hollight.ZIP::'q_17326::type hollight.list |
|
2713 |
=> 'q_17327::type hollight.list |
|
2714 |
=> ('q_17326::type * 'q_17327::type) hollight.list)
|
|
2715 |
(NIL::'q_17326::type hollight.list) |
|
2716 |
(NIL::'q_17327::type hollight.list)) |
|
2717 |
(NIL::('q_17326::type * 'q_17327::type) hollight.list))
|
|
2718 |
((op =::('q_17351::type * 'q_17352::type) hollight.list
|
|
2719 |
=> ('q_17351::type * 'q_17352::type) hollight.list => bool)
|
|
2720 |
((hollight.ZIP::'q_17351::type hollight.list |
|
2721 |
=> 'q_17352::type hollight.list |
|
2722 |
=> ('q_17351::type * 'q_17352::type) hollight.list)
|
|
2723 |
((CONS::'q_17351::type |
|
2724 |
=> 'q_17351::type hollight.list |
|
2725 |
=> 'q_17351::type hollight.list) |
|
2726 |
(h1::'q_17351::type) (t1::'q_17351::type hollight.list)) |
|
2727 |
((CONS::'q_17352::type |
|
2728 |
=> 'q_17352::type hollight.list |
|
2729 |
=> 'q_17352::type hollight.list) |
|
2730 |
(h2::'q_17352::type) (t2::'q_17352::type hollight.list))) |
|
2731 |
((CONS::'q_17351::type * 'q_17352::type |
|
2732 |
=> ('q_17351::type * 'q_17352::type) hollight.list
|
|
2733 |
=> ('q_17351::type * 'q_17352::type) hollight.list)
|
|
2734 |
((Pair::'q_17351::type |
|
2735 |
=> 'q_17352::type => 'q_17351::type * 'q_17352::type) |
|
| 17644 | 2736 |
h1 h2) |
| 19093 | 2737 |
((hollight.ZIP::'q_17351::type hollight.list |
2738 |
=> 'q_17352::type hollight.list |
|
2739 |
=> ('q_17351::type * 'q_17352::type) hollight.list)
|
|
| 17644 | 2740 |
t1 t2)))" |
2741 |
by (import hollight ZIP) |
|
2742 |
||
2743 |
lemma NOT_CONS_NIL: "ALL (x::'A::type) xa::'A::type hollight.list. CONS x xa ~= NIL" |
|
2744 |
by (import hollight NOT_CONS_NIL) |
|
2745 |
||
2746 |
lemma LAST_CLAUSES: "LAST (CONS (h::'A::type) NIL) = h & |
|
2747 |
LAST (CONS h (CONS (k::'A::type) (t::'A::type hollight.list))) = |
|
2748 |
LAST (CONS k t)" |
|
2749 |
by (import hollight LAST_CLAUSES) |
|
2750 |
||
2751 |
lemma APPEND_NIL: "ALL l::'A::type hollight.list. APPEND l NIL = l" |
|
2752 |
by (import hollight APPEND_NIL) |
|
2753 |
||
2754 |
lemma APPEND_ASSOC: "ALL (l::'A::type hollight.list) (m::'A::type hollight.list) |
|
2755 |
n::'A::type hollight.list. APPEND l (APPEND m n) = APPEND (APPEND l m) n" |
|
2756 |
by (import hollight APPEND_ASSOC) |
|
2757 |
||
2758 |
lemma REVERSE_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list. |
|
2759 |
REVERSE (APPEND l m) = APPEND (REVERSE m) (REVERSE l)" |
|
2760 |
by (import hollight REVERSE_APPEND) |
|
2761 |
||
2762 |
lemma REVERSE_REVERSE: "ALL l::'A::type hollight.list. REVERSE (REVERSE l) = l" |
|
2763 |
by (import hollight REVERSE_REVERSE) |
|
2764 |
||
2765 |
lemma CONS_11: "ALL (x::'A::type) (xa::'A::type) (xb::'A::type hollight.list) |
|
2766 |
xc::'A::type hollight.list. (CONS x xb = CONS xa xc) = (x = xa & xb = xc)" |
|
2767 |
by (import hollight CONS_11) |
|
2768 |
||
2769 |
lemma list_CASES: "ALL l::'A::type hollight.list. |
|
2770 |
l = NIL | (EX (h::'A::type) t::'A::type hollight.list. l = CONS h t)" |
|
2771 |
by (import hollight list_CASES) |
|
2772 |
||
2773 |
lemma LENGTH_APPEND: "ALL (l::'A::type hollight.list) m::'A::type hollight.list. |
|
2774 |
LENGTH (APPEND l m) = LENGTH l + LENGTH m" |
|
2775 |
by (import hollight LENGTH_APPEND) |
|
2776 |
||
2777 |
lemma MAP_APPEND: "ALL (f::'A::type => 'B::type) (l1::'A::type hollight.list) |
|
2778 |
l2::'A::type hollight.list. |
|
2779 |
MAP f (APPEND l1 l2) = APPEND (MAP f l1) (MAP f l2)" |
|
2780 |
by (import hollight MAP_APPEND) |
|
2781 |
||
2782 |
lemma LENGTH_MAP: "ALL (l::'A::type hollight.list) f::'A::type => 'B::type. |
|
2783 |
LENGTH (MAP f l) = LENGTH l" |
|
2784 |
by (import hollight LENGTH_MAP) |
|
2785 |
||
| 17652 | 2786 |
lemma LENGTH_EQ_NIL: "ALL l::'A::type hollight.list. (LENGTH l = 0) = (l = NIL)" |
| 17644 | 2787 |
by (import hollight LENGTH_EQ_NIL) |
2788 |
||
| 19093 | 2789 |
lemma LENGTH_EQ_CONS: "ALL (l::'q_17659::type hollight.list) n::nat. |
| 17644 | 2790 |
(LENGTH l = Suc n) = |
| 19093 | 2791 |
(EX (h::'q_17659::type) t::'q_17659::type hollight.list. |
| 17644 | 2792 |
l = CONS h t & LENGTH t = n)" |
2793 |
by (import hollight LENGTH_EQ_CONS) |
|
2794 |
||
2795 |
lemma MAP_o: "ALL (f::'A::type => 'B::type) (g::'B::type => 'C::type) |
|
2796 |
l::'A::type hollight.list. MAP (g o f) l = MAP g (MAP f l)" |
|
2797 |
by (import hollight MAP_o) |
|
2798 |
||
| 19093 | 2799 |
lemma MAP_EQ: "ALL (f::'q_17723::type => 'q_17734::type) |
2800 |
(g::'q_17723::type => 'q_17734::type) l::'q_17723::type hollight.list. |
|
2801 |
ALL_list (%x::'q_17723::type. f x = g x) l --> MAP f l = MAP g l" |
|
| 17644 | 2802 |
by (import hollight MAP_EQ) |
2803 |
||
| 19093 | 2804 |
lemma ALL_IMP: "ALL (P::'q_17764::type => bool) (Q::'q_17764::type => bool) |
2805 |
l::'q_17764::type hollight.list. |
|
2806 |
(ALL x::'q_17764::type. MEM x l & P x --> Q x) & ALL_list P l --> |
|
| 17644 | 2807 |
ALL_list Q l" |
2808 |
by (import hollight ALL_IMP) |
|
2809 |
||
| 19093 | 2810 |
lemma NOT_EX: "ALL (P::'q_17792::type => bool) l::'q_17792::type hollight.list. |
2811 |
(~ EX P l) = ALL_list (%x::'q_17792::type. ~ P x) l" |
|
| 17644 | 2812 |
by (import hollight NOT_EX) |
2813 |
||
| 19093 | 2814 |
lemma NOT_ALL: "ALL (P::'q_17814::type => bool) l::'q_17814::type hollight.list. |
2815 |
(~ ALL_list P l) = EX (%x::'q_17814::type. ~ P x) l" |
|
| 17644 | 2816 |
by (import hollight NOT_ALL) |
2817 |
||
| 19093 | 2818 |
lemma ALL_MAP: "ALL (P::'q_17836::type => bool) (f::'q_17835::type => 'q_17836::type) |
2819 |
l::'q_17835::type hollight.list. |
|
| 17644 | 2820 |
ALL_list P (MAP f l) = ALL_list (P o f) l" |
2821 |
by (import hollight ALL_MAP) |
|
2822 |
||
| 19093 | 2823 |
lemma ALL_T: "All (ALL_list (%x::'q_17854::type. True))" |
| 17644 | 2824 |
by (import hollight ALL_T) |
2825 |
||
| 19093 | 2826 |
lemma MAP_EQ_ALL2: "ALL (l::'q_17879::type hollight.list) m::'q_17879::type hollight.list. |
| 17644 | 2827 |
ALL2 |
| 19093 | 2828 |
(%(x::'q_17879::type) y::'q_17879::type. |
2829 |
(f::'q_17879::type => 'q_17890::type) x = f y) |
|
| 17644 | 2830 |
l m --> |
2831 |
MAP f l = MAP f m" |
|
2832 |
by (import hollight MAP_EQ_ALL2) |
|
2833 |
||
| 19093 | 2834 |
lemma ALL2_MAP: "ALL (P::'q_17921::type => 'q_17922::type => bool) |
2835 |
(f::'q_17922::type => 'q_17921::type) l::'q_17922::type hollight.list. |
|
2836 |
ALL2 P (MAP f l) l = ALL_list (%a::'q_17922::type. P (f a) a) l" |
|
| 17644 | 2837 |
by (import hollight ALL2_MAP) |
2838 |
||
| 19093 | 2839 |
lemma MAP_EQ_DEGEN: "ALL (l::'q_17939::type hollight.list) f::'q_17939::type => 'q_17939::type. |
2840 |
ALL_list (%x::'q_17939::type. f x = x) l --> MAP f l = l" |
|
| 17644 | 2841 |
by (import hollight MAP_EQ_DEGEN) |
2842 |
||
| 19093 | 2843 |
lemma ALL2_AND_RIGHT: "ALL (l::'q_17982::type hollight.list) (m::'q_17981::type hollight.list) |
2844 |
(P::'q_17982::type => bool) Q::'q_17982::type => 'q_17981::type => bool. |
|
2845 |
ALL2 (%(x::'q_17982::type) y::'q_17981::type. P x & Q x y) l m = |
|
| 17644 | 2846 |
(ALL_list P l & ALL2 Q l m)" |
2847 |
by (import hollight ALL2_AND_RIGHT) |
|
2848 |
||
| 19093 | 2849 |
lemma ITLIST_EXTRA: "ALL l::'q_18019::type hollight.list. |
2850 |
ITLIST (f::'q_18019::type => 'q_18018::type => 'q_18018::type) |
|
2851 |
(APPEND l (CONS (a::'q_18019::type) NIL)) (b::'q_18018::type) = |
|
| 17644 | 2852 |
ITLIST f l (f a b)" |
2853 |
by (import hollight ITLIST_EXTRA) |
|
2854 |
||
| 19093 | 2855 |
lemma ALL_MP: "ALL (P::'q_18045::type => bool) (Q::'q_18045::type => bool) |
2856 |
l::'q_18045::type hollight.list. |
|
2857 |
ALL_list (%x::'q_18045::type. P x --> Q x) l & ALL_list P l --> |
|
| 17644 | 2858 |
ALL_list Q l" |
2859 |
by (import hollight ALL_MP) |
|
2860 |
||
| 19093 | 2861 |
lemma AND_ALL: "ALL x::'q_18075::type hollight.list. |
2862 |
(ALL_list (P::'q_18075::type => bool) x & |
|
2863 |
ALL_list (Q::'q_18075::type => bool) x) = |
|
2864 |
ALL_list (%x::'q_18075::type. P x & Q x) x" |
|
| 17644 | 2865 |
by (import hollight AND_ALL) |
2866 |
||
| 19093 | 2867 |
lemma EX_IMP: "ALL (P::'q_18105::type => bool) (Q::'q_18105::type => bool) |
2868 |
l::'q_18105::type hollight.list. |
|
2869 |
(ALL x::'q_18105::type. MEM x l & P x --> Q x) & EX P l --> EX Q l" |
|
| 17644 | 2870 |
by (import hollight EX_IMP) |
2871 |
||
| 19093 | 2872 |
lemma ALL_MEM: "ALL (P::'q_18132::type => bool) l::'q_18132::type hollight.list. |
2873 |
(ALL x::'q_18132::type. MEM x l --> P x) = ALL_list P l" |
|
| 17644 | 2874 |
by (import hollight ALL_MEM) |
2875 |
||
| 19093 | 2876 |
lemma LENGTH_REPLICATE: "ALL (n::nat) x::'q_18150::type. LENGTH (REPLICATE n x) = n" |
| 17644 | 2877 |
by (import hollight LENGTH_REPLICATE) |
2878 |
||
| 19093 | 2879 |
lemma EX_MAP: "ALL (P::'q_18174::type => bool) (f::'q_18173::type => 'q_18174::type) |
2880 |
l::'q_18173::type hollight.list. EX P (MAP f l) = EX (P o f) l" |
|
| 17644 | 2881 |
by (import hollight EX_MAP) |
2882 |
||
| 19093 | 2883 |
lemma EXISTS_EX: "ALL (P::'q_18212::type => 'q_18211::type => bool) |
2884 |
l::'q_18211::type hollight.list. |
|
2885 |
(EX x::'q_18212::type. EX (P x) l) = |
|
2886 |
EX (%s::'q_18211::type. EX x::'q_18212::type. P x s) l" |
|
| 17644 | 2887 |
by (import hollight EXISTS_EX) |
2888 |
||
| 19093 | 2889 |
lemma FORALL_ALL: "ALL (P::'q_18242::type => 'q_18241::type => bool) |
2890 |
l::'q_18241::type hollight.list. |
|
2891 |
(ALL x::'q_18242::type. ALL_list (P x) l) = |
|
2892 |
ALL_list (%s::'q_18241::type. ALL x::'q_18242::type. P x s) l" |
|
| 17644 | 2893 |
by (import hollight FORALL_ALL) |
2894 |
||
| 19093 | 2895 |
lemma MEM_APPEND: "ALL (x::'q_18270::type) (l1::'q_18270::type hollight.list) |
2896 |
l2::'q_18270::type hollight.list. |
|
| 17644 | 2897 |
MEM x (APPEND l1 l2) = (MEM x l1 | MEM x l2)" |
2898 |
by (import hollight MEM_APPEND) |
|
2899 |
||
| 19093 | 2900 |
lemma MEM_MAP: "ALL (f::'q_18306::type => 'q_18303::type) (y::'q_18303::type) |
2901 |
l::'q_18306::type hollight.list. |
|
2902 |
MEM y (MAP f l) = (EX x::'q_18306::type. MEM x l & y = f x)" |
|
| 17644 | 2903 |
by (import hollight MEM_MAP) |
2904 |
||
| 19093 | 2905 |
lemma FILTER_APPEND: "ALL (P::'q_18337::type => bool) (l1::'q_18337::type hollight.list) |
2906 |
l2::'q_18337::type hollight.list. |
|
| 17644 | 2907 |
FILTER P (APPEND l1 l2) = APPEND (FILTER P l1) (FILTER P l2)" |
2908 |
by (import hollight FILTER_APPEND) |
|
2909 |
||
| 19093 | 2910 |
lemma FILTER_MAP: "ALL (P::'q_18364::type => bool) (f::'q_18371::type => 'q_18364::type) |
2911 |
l::'q_18371::type hollight.list. |
|
| 17644 | 2912 |
FILTER P (MAP f l) = MAP f (FILTER (P o f) l)" |
2913 |
by (import hollight FILTER_MAP) |
|
2914 |
||
| 19093 | 2915 |
lemma MEM_FILTER: "ALL (P::'q_18398::type => bool) (l::'q_18398::type hollight.list) |
2916 |
x::'q_18398::type. MEM x (FILTER P l) = (P x & MEM x l)" |
|
| 17644 | 2917 |
by (import hollight MEM_FILTER) |
2918 |
||
| 19093 | 2919 |
lemma EX_MEM: "ALL (P::'q_18419::type => bool) l::'q_18419::type hollight.list. |
2920 |
(EX x::'q_18419::type. P x & MEM x l) = EX P l" |
|
| 17644 | 2921 |
by (import hollight EX_MEM) |
2922 |
||
| 19093 | 2923 |
lemma MAP_FST_ZIP: "ALL (l1::'q_18439::type hollight.list) l2::'q_18441::type hollight.list. |
| 17644 | 2924 |
LENGTH l1 = LENGTH l2 --> MAP fst (hollight.ZIP l1 l2) = l1" |
2925 |
by (import hollight MAP_FST_ZIP) |
|
2926 |
||
| 19093 | 2927 |
lemma MAP_SND_ZIP: "ALL (l1::'q_18470::type hollight.list) l2::'q_18472::type hollight.list. |
| 17644 | 2928 |
LENGTH l1 = LENGTH l2 --> MAP snd (hollight.ZIP l1 l2) = l2" |
2929 |
by (import hollight MAP_SND_ZIP) |
|
2930 |
||
| 19093 | 2931 |
lemma MEM_ASSOC: "ALL (l::('q_18516::type * 'q_18500::type) hollight.list) x::'q_18516::type.
|
| 17644 | 2932 |
MEM (x, ASSOC x l) l = MEM x (MAP fst l)" |
2933 |
by (import hollight MEM_ASSOC) |
|
2934 |
||
| 19093 | 2935 |
lemma ALL_APPEND: "ALL (P::'q_18537::type => bool) (l1::'q_18537::type hollight.list) |
2936 |
l2::'q_18537::type hollight.list. |
|
| 17644 | 2937 |
ALL_list P (APPEND l1 l2) = (ALL_list P l1 & ALL_list P l2)" |
2938 |
by (import hollight ALL_APPEND) |
|
2939 |
||
| 19093 | 2940 |
lemma MEM_EL: "ALL (l::'q_18560::type hollight.list) n::nat. |
| 17644 | 2941 |
< n (LENGTH l) --> MEM (EL n l) l" |
2942 |
by (import hollight MEM_EL) |
|
2943 |
||
| 19093 | 2944 |
lemma ALL2_MAP2: "ALL (l::'q_18603::type hollight.list) m::'q_18604::type hollight.list. |
2945 |
ALL2 (P::'q_18602::type => 'q_18601::type => bool) |
|
2946 |
(MAP (f::'q_18603::type => 'q_18602::type) l) |
|
2947 |
(MAP (g::'q_18604::type => 'q_18601::type) m) = |
|
2948 |
ALL2 (%(x::'q_18603::type) y::'q_18604::type. P (f x) (g y)) l m" |
|
| 17644 | 2949 |
by (import hollight ALL2_MAP2) |
2950 |
||
| 19093 | 2951 |
lemma AND_ALL2: "ALL (P::'q_18650::type => 'q_18649::type => bool) |
2952 |
(Q::'q_18650::type => 'q_18649::type => bool) |
|
2953 |
(x::'q_18650::type hollight.list) xa::'q_18649::type hollight.list. |
|
| 17644 | 2954 |
(ALL2 P x xa & ALL2 Q x xa) = |
| 19093 | 2955 |
ALL2 (%(x::'q_18650::type) y::'q_18649::type. P x y & Q x y) x xa" |
| 17644 | 2956 |
by (import hollight AND_ALL2) |
2957 |
||
| 19093 | 2958 |
lemma ALL2_ALL: "ALL (P::'q_18672::type => 'q_18672::type => bool) |
2959 |
l::'q_18672::type hollight.list. |
|
2960 |
ALL2 P l l = ALL_list (%x::'q_18672::type. P x x) l" |
|
| 17644 | 2961 |
by (import hollight ALL2_ALL) |
2962 |
||
| 19093 | 2963 |
lemma APPEND_EQ_NIL: "ALL (x::'q_18701::type hollight.list) xa::'q_18701::type hollight.list. |
| 17644 | 2964 |
(APPEND x xa = NIL) = (x = NIL & xa = NIL)" |
2965 |
by (import hollight APPEND_EQ_NIL) |
|
2966 |
||
| 19093 | 2967 |
lemma LENGTH_MAP2: "ALL (f::'q_18721::type => 'q_18723::type => 'q_18734::type) |
2968 |
(l::'q_18721::type hollight.list) m::'q_18723::type hollight.list. |
|
| 17644 | 2969 |
LENGTH l = LENGTH m --> LENGTH (MAP2 f l m) = LENGTH m" |
2970 |
by (import hollight LENGTH_MAP2) |
|
2971 |
||
2972 |
lemma MONO_ALL: "(ALL x::'A::type. (P::'A::type => bool) x --> (Q::'A::type => bool) x) --> |
|
2973 |
ALL_list P (l::'A::type hollight.list) --> ALL_list Q l" |
|
2974 |
by (import hollight MONO_ALL) |
|
2975 |
||
2976 |
lemma MONO_ALL2: "(ALL (x::'A::type) y::'B::type. |
|
2977 |
(P::'A::type => 'B::type => bool) x y --> |
|
2978 |
(Q::'A::type => 'B::type => bool) x y) --> |
|
2979 |
ALL2 P (l::'A::type hollight.list) (l'::'B::type hollight.list) --> |
|
2980 |
ALL2 Q l l'" |
|
2981 |
by (import hollight MONO_ALL2) |
|
2982 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
2983 |
definition dist :: "nat * nat => nat" where |
| 17644 | 2984 |
"dist == %u::nat * nat. fst u - snd u + (snd u - fst u)" |
2985 |
||
2986 |
lemma DEF_dist: "dist = (%u::nat * nat. fst u - snd u + (snd u - fst u))" |
|
2987 |
by (import hollight DEF_dist) |
|
2988 |
||
| 17652 | 2989 |
lemma DIST_REFL: "ALL x::nat. dist (x, x) = 0" |
| 17644 | 2990 |
by (import hollight DIST_REFL) |
2991 |
||
| 17652 | 2992 |
lemma DIST_LZERO: "ALL x::nat. dist (0, x) = x" |
| 17644 | 2993 |
by (import hollight DIST_LZERO) |
2994 |
||
| 17652 | 2995 |
lemma DIST_RZERO: "ALL x::nat. dist (x, 0) = x" |
| 17644 | 2996 |
by (import hollight DIST_RZERO) |
2997 |
||
2998 |
lemma DIST_SYM: "ALL (x::nat) xa::nat. dist (x, xa) = dist (xa, x)" |
|
2999 |
by (import hollight DIST_SYM) |
|
3000 |
||
3001 |
lemma DIST_LADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xb, x + xa) = dist (xb, xa)" |
|
3002 |
by (import hollight DIST_LADD) |
|
3003 |
||
3004 |
lemma DIST_RADD: "ALL (x::nat) (xa::nat) xb::nat. dist (x + xa, xb + xa) = dist (x, xb)" |
|
3005 |
by (import hollight DIST_RADD) |
|
3006 |
||
3007 |
lemma DIST_LADD_0: "ALL (x::nat) xa::nat. dist (x + xa, x) = xa" |
|
3008 |
by (import hollight DIST_LADD_0) |
|
3009 |
||
3010 |
lemma DIST_RADD_0: "ALL (x::nat) xa::nat. dist (x, x + xa) = xa" |
|
3011 |
by (import hollight DIST_RADD_0) |
|
3012 |
||
3013 |
lemma DIST_LMUL: "ALL (x::nat) (xa::nat) xb::nat. x * dist (xa, xb) = dist (x * xa, x * xb)" |
|
3014 |
by (import hollight DIST_LMUL) |
|
3015 |
||
3016 |
lemma DIST_RMUL: "ALL (x::nat) (xa::nat) xb::nat. dist (x, xa) * xb = dist (x * xb, xa * xb)" |
|
3017 |
by (import hollight DIST_RMUL) |
|
3018 |
||
| 17652 | 3019 |
lemma DIST_EQ_0: "ALL (x::nat) xa::nat. (dist (x, xa) = 0) = (x = xa)" |
| 17644 | 3020 |
by (import hollight DIST_EQ_0) |
3021 |
||
3022 |
lemma DIST_ELIM_THM: "(P::nat => bool) (dist (x::nat, y::nat)) = |
|
3023 |
(ALL d::nat. (x = y + d --> P d) & (y = x + d --> P d))" |
|
3024 |
by (import hollight DIST_ELIM_THM) |
|
3025 |
||
3026 |
lemma DIST_LE_CASES: "ALL (m::nat) (n::nat) p::nat. |
|
3027 |
<= (dist (m, n)) p = (<= m (n + p) & <= n (m + p))" |
|
3028 |
by (import hollight DIST_LE_CASES) |
|
3029 |
||
3030 |
lemma DIST_ADDBOUND: "ALL (m::nat) n::nat. <= (dist (m, n)) (m + n)" |
|
3031 |
by (import hollight DIST_ADDBOUND) |
|
3032 |
||
3033 |
lemma DIST_TRIANGLE: "ALL (m::nat) (n::nat) p::nat. <= (dist (m, p)) (dist (m, n) + dist (n, p))" |
|
3034 |
by (import hollight DIST_TRIANGLE) |
|
3035 |
||
3036 |
lemma DIST_ADD2: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
3037 |
<= (dist (m + n, p + q)) (dist (m, p) + dist (n, q))" |
|
3038 |
by (import hollight DIST_ADD2) |
|
3039 |
||
3040 |
lemma DIST_ADD2_REV: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
3041 |
<= (dist (m, p)) (dist (m + n, p + q) + dist (n, q))" |
|
3042 |
by (import hollight DIST_ADD2_REV) |
|
3043 |
||
3044 |
lemma DIST_TRIANGLE_LE: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
3045 |
<= (dist (m, n) + dist (n, p)) q --> <= (dist (m, p)) q" |
|
3046 |
by (import hollight DIST_TRIANGLE_LE) |
|
3047 |
||
3048 |
lemma DIST_TRIANGLES_LE: "ALL (m::nat) (n::nat) (p::nat) (q::nat) (r::nat) s::nat. |
|
3049 |
<= (dist (m, n)) r & <= (dist (p, q)) s --> |
|
3050 |
<= (dist (m, p)) (dist (n, q) + (r + s))" |
|
3051 |
by (import hollight DIST_TRIANGLES_LE) |
|
3052 |
||
3053 |
lemma BOUNDS_LINEAR: "ALL (A::nat) (B::nat) C::nat. (ALL n::nat. <= (A * n) (B * n + C)) = <= A B" |
|
3054 |
by (import hollight BOUNDS_LINEAR) |
|
3055 |
||
| 17652 | 3056 |
lemma BOUNDS_LINEAR_0: "ALL (A::nat) B::nat. (ALL n::nat. <= (A * n) B) = (A = 0)" |
| 17644 | 3057 |
by (import hollight BOUNDS_LINEAR_0) |
3058 |
||
3059 |
lemma BOUNDS_DIVIDED: "ALL P::nat => nat. |
|
3060 |
(EX B::nat. ALL n::nat. <= (P n) B) = |
|
3061 |
(EX (x::nat) B::nat. ALL n::nat. <= (n * P n) (x * n + B))" |
|
3062 |
by (import hollight BOUNDS_DIVIDED) |
|
3063 |
||
3064 |
lemma BOUNDS_NOTZERO: "ALL (P::nat => nat => nat) (A::nat) B::nat. |
|
| 17652 | 3065 |
P 0 0 = 0 & (ALL (m::nat) n::nat. <= (P m n) (A * (m + n) + B)) --> |
| 17644 | 3066 |
(EX x::nat. ALL (m::nat) n::nat. <= (P m n) (x * (m + n)))" |
3067 |
by (import hollight BOUNDS_NOTZERO) |
|
3068 |
||
3069 |
lemma BOUNDS_IGNORE: "ALL (P::nat => nat) Q::nat => nat. |
|
3070 |
(EX B::nat. ALL i::nat. <= (P i) (Q i + B)) = |
|
3071 |
(EX (x::nat) N::nat. ALL i::nat. <= N i --> <= (P i) (Q i + x))" |
|
3072 |
by (import hollight BOUNDS_IGNORE) |
|
3073 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3074 |
definition is_nadd :: "(nat => nat) => bool" where |
| 17644 | 3075 |
"is_nadd == |
3076 |
%u::nat => nat. |
|
3077 |
EX B::nat. |
|
3078 |
ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n))" |
|
3079 |
||
3080 |
lemma DEF_is_nadd: "is_nadd = |
|
3081 |
(%u::nat => nat. |
|
3082 |
EX B::nat. |
|
3083 |
ALL (m::nat) n::nat. <= (dist (m * u n, n * u m)) (B * (m + n)))" |
|
3084 |
by (import hollight DEF_is_nadd) |
|
3085 |
||
| 17652 | 3086 |
lemma is_nadd_0: "is_nadd (%n::nat. 0)" |
| 17644 | 3087 |
by (import hollight is_nadd_0) |
3088 |
||
3089 |
typedef (open) nadd = "Collect is_nadd" morphisms "dest_nadd" "mk_nadd" |
|
| 17652 | 3090 |
apply (rule light_ex_imp_nonempty[where t="%n::nat. NUMERAL 0"]) |
| 17644 | 3091 |
by (import hollight TYDEF_nadd) |
3092 |
||
3093 |
syntax |
|
3094 |
dest_nadd :: _ |
|
3095 |
||
3096 |
syntax |
|
3097 |
mk_nadd :: _ |
|
3098 |
||
3099 |
lemmas "TYDEF_nadd_@intern" = typedef_hol2hollight |
|
3100 |
[where a="a :: nadd" and r=r , |
|
3101 |
OF type_definition_nadd] |
|
3102 |
||
3103 |
lemma NADD_CAUCHY: "ALL x::nadd. |
|
3104 |
EX xa::nat. |
|
3105 |
ALL (xb::nat) xc::nat. |
|
3106 |
<= (dist (xb * dest_nadd x xc, xc * dest_nadd x xb)) |
|
3107 |
(xa * (xb + xc))" |
|
3108 |
by (import hollight NADD_CAUCHY) |
|
3109 |
||
3110 |
lemma NADD_BOUND: "ALL x::nadd. |
|
3111 |
EX (xa::nat) B::nat. ALL n::nat. <= (dest_nadd x n) (xa * n + B)" |
|
3112 |
by (import hollight NADD_BOUND) |
|
3113 |
||
3114 |
lemma NADD_MULTIPLICATIVE: "ALL x::nadd. |
|
3115 |
EX xa::nat. |
|
3116 |
ALL (m::nat) n::nat. |
|
3117 |
<= (dist (dest_nadd x (m * n), m * dest_nadd x n)) (xa * m + xa)" |
|
3118 |
by (import hollight NADD_MULTIPLICATIVE) |
|
3119 |
||
3120 |
lemma NADD_ADDITIVE: "ALL x::nadd. |
|
3121 |
EX xa::nat. |
|
3122 |
ALL (m::nat) n::nat. |
|
3123 |
<= (dist (dest_nadd x (m + n), dest_nadd x m + dest_nadd x n)) xa" |
|
3124 |
by (import hollight NADD_ADDITIVE) |
|
3125 |
||
3126 |
lemma NADD_SUC: "ALL x::nadd. |
|
3127 |
EX xa::nat. ALL n::nat. <= (dist (dest_nadd x (Suc n), dest_nadd x n)) xa" |
|
3128 |
by (import hollight NADD_SUC) |
|
3129 |
||
3130 |
lemma NADD_DIST_LEMMA: "ALL x::nadd. |
|
3131 |
EX xa::nat. |
|
3132 |
ALL (m::nat) n::nat. |
|
3133 |
<= (dist (dest_nadd x (m + n), dest_nadd x m)) (xa * n)" |
|
3134 |
by (import hollight NADD_DIST_LEMMA) |
|
3135 |
||
3136 |
lemma NADD_DIST: "ALL x::nadd. |
|
3137 |
EX xa::nat. |
|
3138 |
ALL (m::nat) n::nat. |
|
3139 |
<= (dist (dest_nadd x m, dest_nadd x n)) (xa * dist (m, n))" |
|
3140 |
by (import hollight NADD_DIST) |
|
3141 |
||
3142 |
lemma NADD_ALTMUL: "ALL (x::nadd) y::nadd. |
|
3143 |
EX (A::nat) B::nat. |
|
3144 |
ALL n::nat. |
|
3145 |
<= (dist |
|
3146 |
(n * dest_nadd x (dest_nadd y n), |
|
3147 |
dest_nadd x n * dest_nadd y n)) |
|
3148 |
(A * n + B)" |
|
3149 |
by (import hollight NADD_ALTMUL) |
|
3150 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3151 |
definition nadd_eq :: "nadd => nadd => bool" where |
| 17644 | 3152 |
"nadd_eq == |
3153 |
%(u::nadd) ua::nadd. |
|
3154 |
EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B" |
|
3155 |
||
3156 |
lemma DEF_nadd_eq: "nadd_eq = |
|
3157 |
(%(u::nadd) ua::nadd. |
|
3158 |
EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B)" |
|
3159 |
by (import hollight DEF_nadd_eq) |
|
3160 |
||
3161 |
lemma NADD_EQ_REFL: "ALL x::nadd. nadd_eq x x" |
|
3162 |
by (import hollight NADD_EQ_REFL) |
|
3163 |
||
3164 |
lemma NADD_EQ_SYM: "ALL (x::nadd) y::nadd. nadd_eq x y = nadd_eq y x" |
|
3165 |
by (import hollight NADD_EQ_SYM) |
|
3166 |
||
3167 |
lemma NADD_EQ_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_eq x y & nadd_eq y z --> nadd_eq x z" |
|
3168 |
by (import hollight NADD_EQ_TRANS) |
|
3169 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3170 |
definition nadd_of_num :: "nat => nadd" where |
| 17644 | 3171 |
"nadd_of_num == %u::nat. mk_nadd (op * u)" |
3172 |
||
3173 |
lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. mk_nadd (op * u))" |
|
3174 |
by (import hollight DEF_nadd_of_num) |
|
3175 |
||
3176 |
lemma NADD_OF_NUM: "ALL x::nat. dest_nadd (nadd_of_num x) = op * x" |
|
3177 |
by (import hollight NADD_OF_NUM) |
|
3178 |
||
3179 |
lemma NADD_OF_NUM_WELLDEF: "ALL (m::nat) n::nat. m = n --> nadd_eq (nadd_of_num m) (nadd_of_num n)" |
|
3180 |
by (import hollight NADD_OF_NUM_WELLDEF) |
|
3181 |
||
3182 |
lemma NADD_OF_NUM_EQ: "ALL (m::nat) n::nat. nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)" |
|
3183 |
by (import hollight NADD_OF_NUM_EQ) |
|
3184 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3185 |
definition nadd_le :: "nadd => nadd => bool" where |
| 17644 | 3186 |
"nadd_le == |
3187 |
%(u::nadd) ua::nadd. |
|
3188 |
EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B)" |
|
3189 |
||
3190 |
lemma DEF_nadd_le: "nadd_le = |
|
3191 |
(%(u::nadd) ua::nadd. |
|
3192 |
EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B))" |
|
3193 |
by (import hollight DEF_nadd_le) |
|
3194 |
||
3195 |
lemma NADD_LE_WELLDEF_LEMMA: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. |
|
3196 |
nadd_eq x x' & nadd_eq y y' & nadd_le x y --> nadd_le x' y'" |
|
3197 |
by (import hollight NADD_LE_WELLDEF_LEMMA) |
|
3198 |
||
3199 |
lemma NADD_LE_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. |
|
3200 |
nadd_eq x x' & nadd_eq y y' --> nadd_le x y = nadd_le x' y'" |
|
3201 |
by (import hollight NADD_LE_WELLDEF) |
|
3202 |
||
3203 |
lemma NADD_LE_REFL: "ALL x::nadd. nadd_le x x" |
|
3204 |
by (import hollight NADD_LE_REFL) |
|
3205 |
||
3206 |
lemma NADD_LE_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_le x y & nadd_le y z --> nadd_le x z" |
|
3207 |
by (import hollight NADD_LE_TRANS) |
|
3208 |
||
3209 |
lemma NADD_LE_ANTISYM: "ALL (x::nadd) y::nadd. (nadd_le x y & nadd_le y x) = nadd_eq x y" |
|
3210 |
by (import hollight NADD_LE_ANTISYM) |
|
3211 |
||
3212 |
lemma NADD_LE_TOTAL_LEMMA: "ALL (x::nadd) y::nadd. |
|
3213 |
~ nadd_le x y --> |
|
| 17652 | 3214 |
(ALL B::nat. EX n::nat. n ~= 0 & < (dest_nadd y n + B) (dest_nadd x n))" |
| 17644 | 3215 |
by (import hollight NADD_LE_TOTAL_LEMMA) |
3216 |
||
3217 |
lemma NADD_LE_TOTAL: "ALL (x::nadd) y::nadd. nadd_le x y | nadd_le y x" |
|
3218 |
by (import hollight NADD_LE_TOTAL) |
|
3219 |
||
3220 |
lemma NADD_ARCH: "ALL x::nadd. EX xa::nat. nadd_le x (nadd_of_num xa)" |
|
3221 |
by (import hollight NADD_ARCH) |
|
3222 |
||
3223 |
lemma NADD_OF_NUM_LE: "ALL (m::nat) n::nat. nadd_le (nadd_of_num m) (nadd_of_num n) = <= m n" |
|
3224 |
by (import hollight NADD_OF_NUM_LE) |
|
3225 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3226 |
definition nadd_add :: "nadd => nadd => nadd" where |
| 17644 | 3227 |
"nadd_add == |
3228 |
%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n)" |
|
3229 |
||
3230 |
lemma DEF_nadd_add: "nadd_add = |
|
3231 |
(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n))" |
|
3232 |
by (import hollight DEF_nadd_add) |
|
3233 |
||
3234 |
lemma NADD_ADD: "ALL (x::nadd) y::nadd. |
|
3235 |
dest_nadd (nadd_add x y) = (%n::nat. dest_nadd x n + dest_nadd y n)" |
|
3236 |
by (import hollight NADD_ADD) |
|
3237 |
||
3238 |
lemma NADD_ADD_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. |
|
3239 |
nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_add x y) (nadd_add x' y')" |
|
3240 |
by (import hollight NADD_ADD_WELLDEF) |
|
3241 |
||
3242 |
lemma NADD_ADD_SYM: "ALL (x::nadd) y::nadd. nadd_eq (nadd_add x y) (nadd_add y x)" |
|
3243 |
by (import hollight NADD_ADD_SYM) |
|
3244 |
||
3245 |
lemma NADD_ADD_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3246 |
nadd_eq (nadd_add x (nadd_add y z)) (nadd_add (nadd_add x y) z)" |
|
3247 |
by (import hollight NADD_ADD_ASSOC) |
|
3248 |
||
| 17652 | 3249 |
lemma NADD_ADD_LID: "ALL x::nadd. nadd_eq (nadd_add (nadd_of_num 0) x) x" |
| 17644 | 3250 |
by (import hollight NADD_ADD_LID) |
3251 |
||
3252 |
lemma NADD_ADD_LCANCEL: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3253 |
nadd_eq (nadd_add x y) (nadd_add x z) --> nadd_eq y z" |
|
3254 |
by (import hollight NADD_ADD_LCANCEL) |
|
3255 |
||
3256 |
lemma NADD_LE_ADD: "ALL (x::nadd) y::nadd. nadd_le x (nadd_add x y)" |
|
3257 |
by (import hollight NADD_LE_ADD) |
|
3258 |
||
3259 |
lemma NADD_LE_EXISTS: "ALL (x::nadd) y::nadd. |
|
3260 |
nadd_le x y --> (EX d::nadd. nadd_eq y (nadd_add x d))" |
|
3261 |
by (import hollight NADD_LE_EXISTS) |
|
3262 |
||
3263 |
lemma NADD_OF_NUM_ADD: "ALL (x::nat) xa::nat. |
|
3264 |
nadd_eq (nadd_add (nadd_of_num x) (nadd_of_num xa)) |
|
3265 |
(nadd_of_num (x + xa))" |
|
3266 |
by (import hollight NADD_OF_NUM_ADD) |
|
3267 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3268 |
definition nadd_mul :: "nadd => nadd => nadd" where |
| 17644 | 3269 |
"nadd_mul == |
3270 |
%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n))" |
|
3271 |
||
3272 |
lemma DEF_nadd_mul: "nadd_mul = |
|
3273 |
(%(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n)))" |
|
3274 |
by (import hollight DEF_nadd_mul) |
|
3275 |
||
3276 |
lemma NADD_MUL: "ALL (x::nadd) y::nadd. |
|
3277 |
dest_nadd (nadd_mul x y) = (%n::nat. dest_nadd x (dest_nadd y n))" |
|
3278 |
by (import hollight NADD_MUL) |
|
3279 |
||
3280 |
lemma NADD_MUL_SYM: "ALL (x::nadd) y::nadd. nadd_eq (nadd_mul x y) (nadd_mul y x)" |
|
3281 |
by (import hollight NADD_MUL_SYM) |
|
3282 |
||
3283 |
lemma NADD_MUL_ASSOC: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3284 |
nadd_eq (nadd_mul x (nadd_mul y z)) (nadd_mul (nadd_mul x y) z)" |
|
3285 |
by (import hollight NADD_MUL_ASSOC) |
|
3286 |
||
| 17652 | 3287 |
lemma NADD_MUL_LID: "ALL x::nadd. nadd_eq (nadd_mul (nadd_of_num (NUMERAL_BIT1 0)) x) x" |
| 17644 | 3288 |
by (import hollight NADD_MUL_LID) |
3289 |
||
3290 |
lemma NADD_LDISTRIB: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3291 |
nadd_eq (nadd_mul x (nadd_add y z)) |
|
3292 |
(nadd_add (nadd_mul x y) (nadd_mul x z))" |
|
3293 |
by (import hollight NADD_LDISTRIB) |
|
3294 |
||
3295 |
lemma NADD_MUL_WELLDEF_LEMMA: "ALL (x::nadd) (y::nadd) y'::nadd. |
|
3296 |
nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x y')" |
|
3297 |
by (import hollight NADD_MUL_WELLDEF_LEMMA) |
|
3298 |
||
3299 |
lemma NADD_MUL_WELLDEF: "ALL (x::nadd) (x'::nadd) (y::nadd) y'::nadd. |
|
3300 |
nadd_eq x x' & nadd_eq y y' --> nadd_eq (nadd_mul x y) (nadd_mul x' y')" |
|
3301 |
by (import hollight NADD_MUL_WELLDEF) |
|
3302 |
||
3303 |
lemma NADD_OF_NUM_MUL: "ALL (x::nat) xa::nat. |
|
3304 |
nadd_eq (nadd_mul (nadd_of_num x) (nadd_of_num xa)) |
|
3305 |
(nadd_of_num (x * xa))" |
|
3306 |
by (import hollight NADD_OF_NUM_MUL) |
|
3307 |
||
| 17652 | 3308 |
lemma NADD_LE_0: "All (nadd_le (nadd_of_num 0))" |
| 17644 | 3309 |
by (import hollight NADD_LE_0) |
3310 |
||
3311 |
lemma NADD_EQ_IMP_LE: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_le x y" |
|
3312 |
by (import hollight NADD_EQ_IMP_LE) |
|
3313 |
||
3314 |
lemma NADD_LE_LMUL: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3315 |
nadd_le y z --> nadd_le (nadd_mul x y) (nadd_mul x z)" |
|
3316 |
by (import hollight NADD_LE_LMUL) |
|
3317 |
||
3318 |
lemma NADD_LE_RMUL: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3319 |
nadd_le x y --> nadd_le (nadd_mul x z) (nadd_mul y z)" |
|
3320 |
by (import hollight NADD_LE_RMUL) |
|
3321 |
||
3322 |
lemma NADD_LE_RADD: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3323 |
nadd_le (nadd_add x z) (nadd_add y z) = nadd_le x y" |
|
3324 |
by (import hollight NADD_LE_RADD) |
|
3325 |
||
3326 |
lemma NADD_LE_LADD: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3327 |
nadd_le (nadd_add x y) (nadd_add x z) = nadd_le y z" |
|
3328 |
by (import hollight NADD_LE_LADD) |
|
3329 |
||
3330 |
lemma NADD_RDISTRIB: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3331 |
nadd_eq (nadd_mul (nadd_add x y) z) |
|
3332 |
(nadd_add (nadd_mul x z) (nadd_mul y z))" |
|
3333 |
by (import hollight NADD_RDISTRIB) |
|
3334 |
||
3335 |
lemma NADD_ARCH_MULT: "ALL (x::nadd) k::nat. |
|
| 17652 | 3336 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3337 |
(EX xa::nat. nadd_le (nadd_of_num k) (nadd_mul (nadd_of_num xa) x))" |
3338 |
by (import hollight NADD_ARCH_MULT) |
|
3339 |
||
3340 |
lemma NADD_ARCH_ZERO: "ALL (x::nadd) k::nadd. |
|
3341 |
(ALL n::nat. nadd_le (nadd_mul (nadd_of_num n) x) k) --> |
|
| 17652 | 3342 |
nadd_eq x (nadd_of_num 0)" |
| 17644 | 3343 |
by (import hollight NADD_ARCH_ZERO) |
3344 |
||
3345 |
lemma NADD_ARCH_LEMMA: "ALL (x::nadd) (y::nadd) z::nadd. |
|
3346 |
(ALL n::nat. |
|
3347 |
nadd_le (nadd_mul (nadd_of_num n) x) |
|
3348 |
(nadd_add (nadd_mul (nadd_of_num n) y) z)) --> |
|
3349 |
nadd_le x y" |
|
3350 |
by (import hollight NADD_ARCH_LEMMA) |
|
3351 |
||
3352 |
lemma NADD_COMPLETE: "ALL P::nadd => bool. |
|
3353 |
Ex P & (EX M::nadd. ALL x::nadd. P x --> nadd_le x M) --> |
|
3354 |
(EX M::nadd. |
|
3355 |
(ALL x::nadd. P x --> nadd_le x M) & |
|
3356 |
(ALL M'::nadd. (ALL x::nadd. P x --> nadd_le x M') --> nadd_le M M'))" |
|
3357 |
by (import hollight NADD_COMPLETE) |
|
3358 |
||
3359 |
lemma NADD_UBOUND: "ALL x::nadd. |
|
3360 |
EX (xa::nat) N::nat. ALL n::nat. <= N n --> <= (dest_nadd x n) (xa * n)" |
|
3361 |
by (import hollight NADD_UBOUND) |
|
3362 |
||
3363 |
lemma NADD_NONZERO: "ALL x::nadd. |
|
| 17652 | 3364 |
~ nadd_eq x (nadd_of_num 0) --> |
3365 |
(EX N::nat. ALL n::nat. <= N n --> dest_nadd x n ~= 0)" |
|
| 17644 | 3366 |
by (import hollight NADD_NONZERO) |
3367 |
||
3368 |
lemma NADD_LBOUND: "ALL x::nadd. |
|
| 17652 | 3369 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3370 |
(EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))" |
3371 |
by (import hollight NADD_LBOUND) |
|
3372 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3373 |
definition nadd_rinv :: "nadd => nat => nat" where |
| 17644 | 3374 |
"nadd_rinv == %(u::nadd) n::nat. DIV (n * n) (dest_nadd u n)" |
3375 |
||
3376 |
lemma DEF_nadd_rinv: "nadd_rinv = (%(u::nadd) n::nat. DIV (n * n) (dest_nadd u n))" |
|
3377 |
by (import hollight DEF_nadd_rinv) |
|
3378 |
||
3379 |
lemma NADD_MUL_LINV_LEMMA0: "ALL x::nadd. |
|
| 17652 | 3380 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3381 |
(EX (xa::nat) B::nat. ALL i::nat. <= (nadd_rinv x i) (xa * i + B))" |
3382 |
by (import hollight NADD_MUL_LINV_LEMMA0) |
|
3383 |
||
3384 |
lemma NADD_MUL_LINV_LEMMA1: "ALL (x::nadd) n::nat. |
|
| 17652 | 3385 |
dest_nadd x n ~= 0 --> |
| 17644 | 3386 |
<= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n)" |
3387 |
by (import hollight NADD_MUL_LINV_LEMMA1) |
|
3388 |
||
3389 |
lemma NADD_MUL_LINV_LEMMA2: "ALL x::nadd. |
|
| 17652 | 3390 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3391 |
(EX N::nat. |
3392 |
ALL n::nat. |
|
3393 |
<= N n --> |
|
3394 |
<= (dist (dest_nadd x n * nadd_rinv x n, n * n)) (dest_nadd x n))" |
|
3395 |
by (import hollight NADD_MUL_LINV_LEMMA2) |
|
3396 |
||
3397 |
lemma NADD_MUL_LINV_LEMMA3: "ALL x::nadd. |
|
| 17652 | 3398 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3399 |
(EX N::nat. |
3400 |
ALL (m::nat) n::nat. |
|
3401 |
<= N n --> |
|
3402 |
<= (dist |
|
3403 |
(m * (dest_nadd x m * (dest_nadd x n * nadd_rinv x n)), |
|
3404 |
m * (dest_nadd x m * (n * n)))) |
|
3405 |
(m * (dest_nadd x m * dest_nadd x n)))" |
|
3406 |
by (import hollight NADD_MUL_LINV_LEMMA3) |
|
3407 |
||
3408 |
lemma NADD_MUL_LINV_LEMMA4: "ALL x::nadd. |
|
| 17652 | 3409 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3410 |
(EX N::nat. |
3411 |
ALL (m::nat) n::nat. |
|
3412 |
<= N m & <= N n --> |
|
3413 |
<= (dest_nadd x m * dest_nadd x n * |
|
3414 |
dist (m * nadd_rinv x n, n * nadd_rinv x m)) |
|
3415 |
(m * n * dist (m * dest_nadd x n, n * dest_nadd x m) + |
|
3416 |
dest_nadd x m * dest_nadd x n * (m + n)))" |
|
3417 |
by (import hollight NADD_MUL_LINV_LEMMA4) |
|
3418 |
||
3419 |
lemma NADD_MUL_LINV_LEMMA5: "ALL x::nadd. |
|
| 17652 | 3420 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3421 |
(EX (B::nat) N::nat. |
3422 |
ALL (m::nat) n::nat. |
|
3423 |
<= N m & <= N n --> |
|
3424 |
<= (dest_nadd x m * dest_nadd x n * |
|
3425 |
dist (m * nadd_rinv x n, n * nadd_rinv x m)) |
|
3426 |
(B * (m * n * (m + n))))" |
|
3427 |
by (import hollight NADD_MUL_LINV_LEMMA5) |
|
3428 |
||
3429 |
lemma NADD_MUL_LINV_LEMMA6: "ALL x::nadd. |
|
| 17652 | 3430 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3431 |
(EX (B::nat) N::nat. |
3432 |
ALL (m::nat) n::nat. |
|
3433 |
<= N m & <= N n --> |
|
3434 |
<= (m * n * dist (m * nadd_rinv x n, n * nadd_rinv x m)) |
|
3435 |
(B * (m * n * (m + n))))" |
|
3436 |
by (import hollight NADD_MUL_LINV_LEMMA6) |
|
3437 |
||
3438 |
lemma NADD_MUL_LINV_LEMMA7: "ALL x::nadd. |
|
| 17652 | 3439 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3440 |
(EX (B::nat) N::nat. |
3441 |
ALL (m::nat) n::nat. |
|
3442 |
<= N m & <= N n --> |
|
3443 |
<= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" |
|
3444 |
by (import hollight NADD_MUL_LINV_LEMMA7) |
|
3445 |
||
3446 |
lemma NADD_MUL_LINV_LEMMA7a: "ALL x::nadd. |
|
| 17652 | 3447 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3448 |
(ALL N::nat. |
3449 |
EX (A::nat) B::nat. |
|
3450 |
ALL (m::nat) n::nat. |
|
3451 |
<= m N --> |
|
3452 |
<= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (A * n + B))" |
|
3453 |
by (import hollight NADD_MUL_LINV_LEMMA7a) |
|
3454 |
||
3455 |
lemma NADD_MUL_LINV_LEMMA8: "ALL x::nadd. |
|
| 17652 | 3456 |
~ nadd_eq x (nadd_of_num 0) --> |
| 17644 | 3457 |
(EX B::nat. |
3458 |
ALL (m::nat) n::nat. |
|
3459 |
<= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" |
|
3460 |
by (import hollight NADD_MUL_LINV_LEMMA8) |
|
3461 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3462 |
definition nadd_inv :: "nadd => nadd" where |
| 17644 | 3463 |
"nadd_inv == |
3464 |
%u::nadd. |
|
| 17652 | 3465 |
COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) (mk_nadd (nadd_rinv u))" |
| 17644 | 3466 |
|
3467 |
lemma DEF_nadd_inv: "nadd_inv = |
|
3468 |
(%u::nadd. |
|
| 17652 | 3469 |
COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) |
| 17644 | 3470 |
(mk_nadd (nadd_rinv u)))" |
3471 |
by (import hollight DEF_nadd_inv) |
|
3472 |
||
3473 |
lemma NADD_INV: "ALL x::nadd. |
|
3474 |
dest_nadd (nadd_inv x) = |
|
| 17652 | 3475 |
COND (nadd_eq x (nadd_of_num 0)) (%n::nat. 0) (nadd_rinv x)" |
| 17644 | 3476 |
by (import hollight NADD_INV) |
3477 |
||
3478 |
lemma NADD_MUL_LINV: "ALL x::nadd. |
|
| 17652 | 3479 |
~ nadd_eq x (nadd_of_num 0) --> |
3480 |
nadd_eq (nadd_mul (nadd_inv x) x) (nadd_of_num (NUMERAL_BIT1 0))" |
|
| 17644 | 3481 |
by (import hollight NADD_MUL_LINV) |
3482 |
||
| 17652 | 3483 |
lemma NADD_INV_0: "nadd_eq (nadd_inv (nadd_of_num 0)) (nadd_of_num 0)" |
| 17644 | 3484 |
by (import hollight NADD_INV_0) |
3485 |
||
3486 |
lemma NADD_INV_WELLDEF: "ALL (x::nadd) y::nadd. nadd_eq x y --> nadd_eq (nadd_inv x) (nadd_inv y)" |
|
3487 |
by (import hollight NADD_INV_WELLDEF) |
|
3488 |
||
3489 |
typedef (open) hreal = "{s::nadd => bool. EX x::nadd. s = nadd_eq x}" morphisms "dest_hreal" "mk_hreal"
|
|
3490 |
apply (rule light_ex_imp_nonempty[where t="nadd_eq (x::nadd)"]) |
|
3491 |
by (import hollight TYDEF_hreal) |
|
3492 |
||
3493 |
syntax |
|
3494 |
dest_hreal :: _ |
|
3495 |
||
3496 |
syntax |
|
3497 |
mk_hreal :: _ |
|
3498 |
||
3499 |
lemmas "TYDEF_hreal_@intern" = typedef_hol2hollight |
|
3500 |
[where a="a :: hreal" and r=r , |
|
3501 |
OF type_definition_hreal] |
|
3502 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3503 |
definition hreal_of_num :: "nat => hreal" where |
| 17644 | 3504 |
"hreal_of_num == %m::nat. mk_hreal (nadd_eq (nadd_of_num m))" |
3505 |
||
3506 |
lemma DEF_hreal_of_num: "hreal_of_num = (%m::nat. mk_hreal (nadd_eq (nadd_of_num m)))" |
|
3507 |
by (import hollight DEF_hreal_of_num) |
|
3508 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3509 |
definition hreal_add :: "hreal => hreal => hreal" where |
| 17644 | 3510 |
"hreal_add == |
3511 |
%(x::hreal) y::hreal. |
|
3512 |
mk_hreal |
|
3513 |
(%u::nadd. |
|
3514 |
EX (xa::nadd) ya::nadd. |
|
3515 |
nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya)" |
|
3516 |
||
3517 |
lemma DEF_hreal_add: "hreal_add = |
|
3518 |
(%(x::hreal) y::hreal. |
|
3519 |
mk_hreal |
|
3520 |
(%u::nadd. |
|
3521 |
EX (xa::nadd) ya::nadd. |
|
3522 |
nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya))" |
|
3523 |
by (import hollight DEF_hreal_add) |
|
3524 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3525 |
definition hreal_mul :: "hreal => hreal => hreal" where |
| 17644 | 3526 |
"hreal_mul == |
3527 |
%(x::hreal) y::hreal. |
|
3528 |
mk_hreal |
|
3529 |
(%u::nadd. |
|
3530 |
EX (xa::nadd) ya::nadd. |
|
3531 |
nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya)" |
|
3532 |
||
3533 |
lemma DEF_hreal_mul: "hreal_mul = |
|
3534 |
(%(x::hreal) y::hreal. |
|
3535 |
mk_hreal |
|
3536 |
(%u::nadd. |
|
3537 |
EX (xa::nadd) ya::nadd. |
|
3538 |
nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya))" |
|
3539 |
by (import hollight DEF_hreal_mul) |
|
3540 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3541 |
definition hreal_le :: "hreal => hreal => bool" where |
| 17644 | 3542 |
"hreal_le == |
3543 |
%(x::hreal) y::hreal. |
|
3544 |
SOME u::bool. |
|
3545 |
EX (xa::nadd) ya::nadd. |
|
3546 |
nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya" |
|
3547 |
||
3548 |
lemma DEF_hreal_le: "hreal_le = |
|
3549 |
(%(x::hreal) y::hreal. |
|
3550 |
SOME u::bool. |
|
3551 |
EX (xa::nadd) ya::nadd. |
|
3552 |
nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)" |
|
3553 |
by (import hollight DEF_hreal_le) |
|
3554 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3555 |
definition hreal_inv :: "hreal => hreal" where |
| 17644 | 3556 |
"hreal_inv == |
3557 |
%x::hreal. |
|
3558 |
mk_hreal |
|
3559 |
(%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa)" |
|
3560 |
||
3561 |
lemma DEF_hreal_inv: "hreal_inv = |
|
3562 |
(%x::hreal. |
|
3563 |
mk_hreal |
|
3564 |
(%u::nadd. EX xa::nadd. nadd_eq (nadd_inv xa) u & dest_hreal x xa))" |
|
3565 |
by (import hollight DEF_hreal_inv) |
|
3566 |
||
3567 |
lemma HREAL_LE_EXISTS_DEF: "ALL (m::hreal) n::hreal. hreal_le m n = (EX d::hreal. n = hreal_add m d)" |
|
3568 |
by (import hollight HREAL_LE_EXISTS_DEF) |
|
3569 |
||
3570 |
lemma HREAL_EQ_ADD_LCANCEL: "ALL (m::hreal) (n::hreal) p::hreal. |
|
3571 |
(hreal_add m n = hreal_add m p) = (n = p)" |
|
3572 |
by (import hollight HREAL_EQ_ADD_LCANCEL) |
|
3573 |
||
3574 |
lemma HREAL_EQ_ADD_RCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal. |
|
3575 |
(hreal_add x xb = hreal_add xa xb) = (x = xa)" |
|
3576 |
by (import hollight HREAL_EQ_ADD_RCANCEL) |
|
3577 |
||
3578 |
lemma HREAL_LE_ADD_LCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal. |
|
3579 |
hreal_le (hreal_add x xa) (hreal_add x xb) = hreal_le xa xb" |
|
3580 |
by (import hollight HREAL_LE_ADD_LCANCEL) |
|
3581 |
||
3582 |
lemma HREAL_LE_ADD_RCANCEL: "ALL (x::hreal) (xa::hreal) xb::hreal. |
|
3583 |
hreal_le (hreal_add x xb) (hreal_add xa xb) = hreal_le x xa" |
|
3584 |
by (import hollight HREAL_LE_ADD_RCANCEL) |
|
3585 |
||
| 17652 | 3586 |
lemma HREAL_ADD_RID: "ALL x::hreal. hreal_add x (hreal_of_num 0) = x" |
| 17644 | 3587 |
by (import hollight HREAL_ADD_RID) |
3588 |
||
3589 |
lemma HREAL_ADD_RDISTRIB: "ALL (x::hreal) (xa::hreal) xb::hreal. |
|
3590 |
hreal_mul (hreal_add x xa) xb = |
|
3591 |
hreal_add (hreal_mul x xb) (hreal_mul xa xb)" |
|
3592 |
by (import hollight HREAL_ADD_RDISTRIB) |
|
3593 |
||
| 17652 | 3594 |
lemma HREAL_MUL_LZERO: "ALL m::hreal. hreal_mul (hreal_of_num 0) m = hreal_of_num 0" |
| 17644 | 3595 |
by (import hollight HREAL_MUL_LZERO) |
3596 |
||
| 17652 | 3597 |
lemma HREAL_MUL_RZERO: "ALL x::hreal. hreal_mul x (hreal_of_num 0) = hreal_of_num 0" |
| 17644 | 3598 |
by (import hollight HREAL_MUL_RZERO) |
3599 |
||
3600 |
lemma HREAL_ADD_AC: "hreal_add (m::hreal) (n::hreal) = hreal_add n m & |
|
3601 |
hreal_add (hreal_add m n) (p::hreal) = hreal_add m (hreal_add n p) & |
|
3602 |
hreal_add m (hreal_add n p) = hreal_add n (hreal_add m p)" |
|
3603 |
by (import hollight HREAL_ADD_AC) |
|
3604 |
||
3605 |
lemma HREAL_LE_ADD2: "ALL (a::hreal) (b::hreal) (c::hreal) d::hreal. |
|
3606 |
hreal_le a b & hreal_le c d --> hreal_le (hreal_add a c) (hreal_add b d)" |
|
3607 |
by (import hollight HREAL_LE_ADD2) |
|
3608 |
||
3609 |
lemma HREAL_LE_MUL_RCANCEL_IMP: "ALL (a::hreal) (b::hreal) c::hreal. |
|
3610 |
hreal_le a b --> hreal_le (hreal_mul a c) (hreal_mul b c)" |
|
3611 |
by (import hollight HREAL_LE_MUL_RCANCEL_IMP) |
|
3612 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3613 |
definition treal_of_num :: "nat => hreal * hreal" where |
| 17652 | 3614 |
"treal_of_num == %u::nat. (hreal_of_num u, hreal_of_num 0)" |
3615 |
||
3616 |
lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num 0))" |
|
| 17644 | 3617 |
by (import hollight DEF_treal_of_num) |
3618 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3619 |
definition treal_neg :: "hreal * hreal => hreal * hreal" where |
| 17644 | 3620 |
"treal_neg == %u::hreal * hreal. (snd u, fst u)" |
3621 |
||
3622 |
lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))" |
|
3623 |
by (import hollight DEF_treal_neg) |
|
3624 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3625 |
definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where |
| 17644 | 3626 |
"treal_add == |
3627 |
%(u::hreal * hreal) ua::hreal * hreal. |
|
3628 |
(hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))" |
|
3629 |
||
3630 |
lemma DEF_treal_add: "treal_add = |
|
3631 |
(%(u::hreal * hreal) ua::hreal * hreal. |
|
3632 |
(hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))" |
|
3633 |
by (import hollight DEF_treal_add) |
|
3634 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3635 |
definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where |
| 17644 | 3636 |
"treal_mul == |
3637 |
%(u::hreal * hreal) ua::hreal * hreal. |
|
3638 |
(hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)), |
|
3639 |
hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua)))" |
|
3640 |
||
3641 |
lemma DEF_treal_mul: "treal_mul = |
|
3642 |
(%(u::hreal * hreal) ua::hreal * hreal. |
|
3643 |
(hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)), |
|
3644 |
hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))" |
|
3645 |
by (import hollight DEF_treal_mul) |
|
3646 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3647 |
definition treal_le :: "hreal * hreal => hreal * hreal => bool" where |
| 17644 | 3648 |
"treal_le == |
3649 |
%(u::hreal * hreal) ua::hreal * hreal. |
|
3650 |
hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))" |
|
3651 |
||
3652 |
lemma DEF_treal_le: "treal_le = |
|
3653 |
(%(u::hreal * hreal) ua::hreal * hreal. |
|
3654 |
hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))" |
|
3655 |
by (import hollight DEF_treal_le) |
|
3656 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3657 |
definition treal_inv :: "hreal * hreal => hreal * hreal" where |
| 17644 | 3658 |
"treal_inv == |
3659 |
%u::hreal * hreal. |
|
| 17652 | 3660 |
COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0) |
| 17644 | 3661 |
(COND (hreal_le (snd u) (fst u)) |
3662 |
(hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), |
|
| 17652 | 3663 |
hreal_of_num 0) |
3664 |
(hreal_of_num 0, |
|
| 17644 | 3665 |
hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d)))" |
3666 |
||
3667 |
lemma DEF_treal_inv: "treal_inv = |
|
3668 |
(%u::hreal * hreal. |
|
| 17652 | 3669 |
COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0) |
| 17644 | 3670 |
(COND (hreal_le (snd u) (fst u)) |
3671 |
(hreal_inv (SOME d::hreal. fst u = hreal_add (snd u) d), |
|
| 17652 | 3672 |
hreal_of_num 0) |
3673 |
(hreal_of_num 0, |
|
| 17644 | 3674 |
hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))" |
3675 |
by (import hollight DEF_treal_inv) |
|
3676 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3677 |
definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where |
| 17644 | 3678 |
"treal_eq == |
3679 |
%(u::hreal * hreal) ua::hreal * hreal. |
|
3680 |
hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)" |
|
3681 |
||
3682 |
lemma DEF_treal_eq: "treal_eq = |
|
3683 |
(%(u::hreal * hreal) ua::hreal * hreal. |
|
3684 |
hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u))" |
|
3685 |
by (import hollight DEF_treal_eq) |
|
3686 |
||
3687 |
lemma TREAL_EQ_REFL: "ALL x::hreal * hreal. treal_eq x x" |
|
3688 |
by (import hollight TREAL_EQ_REFL) |
|
3689 |
||
3690 |
lemma TREAL_EQ_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y = treal_eq y x" |
|
3691 |
by (import hollight TREAL_EQ_SYM) |
|
3692 |
||
3693 |
lemma TREAL_EQ_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3694 |
treal_eq x y & treal_eq y z --> treal_eq x z" |
|
3695 |
by (import hollight TREAL_EQ_TRANS) |
|
3696 |
||
3697 |
lemma TREAL_EQ_AP: "ALL (x::hreal * hreal) y::hreal * hreal. x = y --> treal_eq x y" |
|
3698 |
by (import hollight TREAL_EQ_AP) |
|
3699 |
||
3700 |
lemma TREAL_OF_NUM_EQ: "ALL (x::nat) xa::nat. treal_eq (treal_of_num x) (treal_of_num xa) = (x = xa)" |
|
3701 |
by (import hollight TREAL_OF_NUM_EQ) |
|
3702 |
||
3703 |
lemma TREAL_OF_NUM_LE: "ALL (x::nat) xa::nat. treal_le (treal_of_num x) (treal_of_num xa) = <= x xa" |
|
3704 |
by (import hollight TREAL_OF_NUM_LE) |
|
3705 |
||
3706 |
lemma TREAL_OF_NUM_ADD: "ALL (x::nat) xa::nat. |
|
3707 |
treal_eq (treal_add (treal_of_num x) (treal_of_num xa)) |
|
3708 |
(treal_of_num (x + xa))" |
|
3709 |
by (import hollight TREAL_OF_NUM_ADD) |
|
3710 |
||
3711 |
lemma TREAL_OF_NUM_MUL: "ALL (x::nat) xa::nat. |
|
3712 |
treal_eq (treal_mul (treal_of_num x) (treal_of_num xa)) |
|
3713 |
(treal_of_num (x * xa))" |
|
3714 |
by (import hollight TREAL_OF_NUM_MUL) |
|
3715 |
||
3716 |
lemma TREAL_ADD_SYM_EQ: "ALL (x::hreal * hreal) y::hreal * hreal. treal_add x y = treal_add y x" |
|
3717 |
by (import hollight TREAL_ADD_SYM_EQ) |
|
3718 |
||
3719 |
lemma TREAL_MUL_SYM_EQ: "ALL (x::hreal * hreal) y::hreal * hreal. treal_mul x y = treal_mul y x" |
|
3720 |
by (import hollight TREAL_MUL_SYM_EQ) |
|
3721 |
||
3722 |
lemma TREAL_ADD_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. |
|
3723 |
treal_eq (treal_add x y) (treal_add y x)" |
|
3724 |
by (import hollight TREAL_ADD_SYM) |
|
3725 |
||
3726 |
lemma TREAL_ADD_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3727 |
treal_eq (treal_add x (treal_add y z)) (treal_add (treal_add x y) z)" |
|
3728 |
by (import hollight TREAL_ADD_ASSOC) |
|
3729 |
||
| 17652 | 3730 |
lemma TREAL_ADD_LID: "ALL x::hreal * hreal. treal_eq (treal_add (treal_of_num 0) x) x" |
| 17644 | 3731 |
by (import hollight TREAL_ADD_LID) |
3732 |
||
| 17652 | 3733 |
lemma TREAL_ADD_LINV: "ALL x::hreal * hreal. treal_eq (treal_add (treal_neg x) x) (treal_of_num 0)" |
| 17644 | 3734 |
by (import hollight TREAL_ADD_LINV) |
3735 |
||
3736 |
lemma TREAL_MUL_SYM: "ALL (x::hreal * hreal) y::hreal * hreal. |
|
3737 |
treal_eq (treal_mul x y) (treal_mul y x)" |
|
3738 |
by (import hollight TREAL_MUL_SYM) |
|
3739 |
||
3740 |
lemma TREAL_MUL_ASSOC: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3741 |
treal_eq (treal_mul x (treal_mul y z)) (treal_mul (treal_mul x y) z)" |
|
3742 |
by (import hollight TREAL_MUL_ASSOC) |
|
3743 |
||
3744 |
lemma TREAL_MUL_LID: "ALL x::hreal * hreal. |
|
| 17652 | 3745 |
treal_eq (treal_mul (treal_of_num (NUMERAL_BIT1 0)) x) x" |
| 17644 | 3746 |
by (import hollight TREAL_MUL_LID) |
3747 |
||
3748 |
lemma TREAL_ADD_LDISTRIB: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3749 |
treal_eq (treal_mul x (treal_add y z)) |
|
3750 |
(treal_add (treal_mul x y) (treal_mul x z))" |
|
3751 |
by (import hollight TREAL_ADD_LDISTRIB) |
|
3752 |
||
3753 |
lemma TREAL_LE_REFL: "ALL x::hreal * hreal. treal_le x x" |
|
3754 |
by (import hollight TREAL_LE_REFL) |
|
3755 |
||
3756 |
lemma TREAL_LE_ANTISYM: "ALL (x::hreal * hreal) y::hreal * hreal. |
|
3757 |
(treal_le x y & treal_le y x) = treal_eq x y" |
|
3758 |
by (import hollight TREAL_LE_ANTISYM) |
|
3759 |
||
3760 |
lemma TREAL_LE_TRANS: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3761 |
treal_le x y & treal_le y z --> treal_le x z" |
|
3762 |
by (import hollight TREAL_LE_TRANS) |
|
3763 |
||
3764 |
lemma TREAL_LE_TOTAL: "ALL (x::hreal * hreal) y::hreal * hreal. treal_le x y | treal_le y x" |
|
3765 |
by (import hollight TREAL_LE_TOTAL) |
|
3766 |
||
3767 |
lemma TREAL_LE_LADD_IMP: "ALL (x::hreal * hreal) (y::hreal * hreal) z::hreal * hreal. |
|
3768 |
treal_le y z --> treal_le (treal_add x y) (treal_add x z)" |
|
3769 |
by (import hollight TREAL_LE_LADD_IMP) |
|
3770 |
||
3771 |
lemma TREAL_LE_MUL: "ALL (x::hreal * hreal) y::hreal * hreal. |
|
| 17652 | 3772 |
treal_le (treal_of_num 0) x & treal_le (treal_of_num 0) y --> |
3773 |
treal_le (treal_of_num 0) (treal_mul x y)" |
|
| 17644 | 3774 |
by (import hollight TREAL_LE_MUL) |
3775 |
||
| 17652 | 3776 |
lemma TREAL_INV_0: "treal_eq (treal_inv (treal_of_num 0)) (treal_of_num 0)" |
| 17644 | 3777 |
by (import hollight TREAL_INV_0) |
3778 |
||
3779 |
lemma TREAL_MUL_LINV: "ALL x::hreal * hreal. |
|
| 17652 | 3780 |
~ treal_eq x (treal_of_num 0) --> |
3781 |
treal_eq (treal_mul (treal_inv x) x) (treal_of_num (NUMERAL_BIT1 0))" |
|
| 17644 | 3782 |
by (import hollight TREAL_MUL_LINV) |
3783 |
||
3784 |
lemma TREAL_OF_NUM_WELLDEF: "ALL (m::nat) n::nat. m = n --> treal_eq (treal_of_num m) (treal_of_num n)" |
|
3785 |
by (import hollight TREAL_OF_NUM_WELLDEF) |
|
3786 |
||
3787 |
lemma TREAL_NEG_WELLDEF: "ALL (x1::hreal * hreal) x2::hreal * hreal. |
|
3788 |
treal_eq x1 x2 --> treal_eq (treal_neg x1) (treal_neg x2)" |
|
3789 |
by (import hollight TREAL_NEG_WELLDEF) |
|
3790 |
||
3791 |
lemma TREAL_ADD_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal. |
|
3792 |
treal_eq x1 x2 --> treal_eq (treal_add x1 y) (treal_add x2 y)" |
|
3793 |
by (import hollight TREAL_ADD_WELLDEFR) |
|
3794 |
||
3795 |
lemma TREAL_ADD_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) |
|
3796 |
y2::hreal * hreal. |
|
3797 |
treal_eq x1 x2 & treal_eq y1 y2 --> |
|
3798 |
treal_eq (treal_add x1 y1) (treal_add x2 y2)" |
|
3799 |
by (import hollight TREAL_ADD_WELLDEF) |
|
3800 |
||
3801 |
lemma TREAL_MUL_WELLDEFR: "ALL (x1::hreal * hreal) (x2::hreal * hreal) y::hreal * hreal. |
|
3802 |
treal_eq x1 x2 --> treal_eq (treal_mul x1 y) (treal_mul x2 y)" |
|
3803 |
by (import hollight TREAL_MUL_WELLDEFR) |
|
3804 |
||
3805 |
lemma TREAL_MUL_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) |
|
3806 |
y2::hreal * hreal. |
|
3807 |
treal_eq x1 x2 & treal_eq y1 y2 --> |
|
3808 |
treal_eq (treal_mul x1 y1) (treal_mul x2 y2)" |
|
3809 |
by (import hollight TREAL_MUL_WELLDEF) |
|
3810 |
||
3811 |
lemma TREAL_EQ_IMP_LE: "ALL (x::hreal * hreal) y::hreal * hreal. treal_eq x y --> treal_le x y" |
|
3812 |
by (import hollight TREAL_EQ_IMP_LE) |
|
3813 |
||
3814 |
lemma TREAL_LE_WELLDEF: "ALL (x1::hreal * hreal) (x2::hreal * hreal) (y1::hreal * hreal) |
|
3815 |
y2::hreal * hreal. |
|
3816 |
treal_eq x1 x2 & treal_eq y1 y2 --> treal_le x1 y1 = treal_le x2 y2" |
|
3817 |
by (import hollight TREAL_LE_WELLDEF) |
|
3818 |
||
3819 |
lemma TREAL_INV_WELLDEF: "ALL (x::hreal * hreal) y::hreal * hreal. |
|
3820 |
treal_eq x y --> treal_eq (treal_inv x) (treal_inv y)" |
|
3821 |
by (import hollight TREAL_INV_WELLDEF) |
|
3822 |
||
3823 |
typedef (open) real = "{s::hreal * hreal => bool. EX x::hreal * hreal. s = treal_eq x}" morphisms "dest_real" "mk_real"
|
|
3824 |
apply (rule light_ex_imp_nonempty[where t="treal_eq (x::hreal * hreal)"]) |
|
3825 |
by (import hollight TYDEF_real) |
|
3826 |
||
3827 |
syntax |
|
3828 |
dest_real :: _ |
|
3829 |
||
3830 |
syntax |
|
3831 |
mk_real :: _ |
|
3832 |
||
3833 |
lemmas "TYDEF_real_@intern" = typedef_hol2hollight |
|
3834 |
[where a="a :: hollight.real" and r=r , |
|
3835 |
OF type_definition_real] |
|
3836 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3837 |
definition real_of_num :: "nat => hollight.real" where |
| 17644 | 3838 |
"real_of_num == %m::nat. mk_real (treal_eq (treal_of_num m))" |
3839 |
||
3840 |
lemma DEF_real_of_num: "real_of_num = (%m::nat. mk_real (treal_eq (treal_of_num m)))" |
|
3841 |
by (import hollight DEF_real_of_num) |
|
3842 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3843 |
definition real_neg :: "hollight.real => hollight.real" where |
| 17644 | 3844 |
"real_neg == |
3845 |
%x1::hollight.real. |
|
3846 |
mk_real |
|
3847 |
(%u::hreal * hreal. |
|
3848 |
EX x1a::hreal * hreal. |
|
3849 |
treal_eq (treal_neg x1a) u & dest_real x1 x1a)" |
|
3850 |
||
3851 |
lemma DEF_real_neg: "real_neg = |
|
3852 |
(%x1::hollight.real. |
|
3853 |
mk_real |
|
3854 |
(%u::hreal * hreal. |
|
3855 |
EX x1a::hreal * hreal. |
|
3856 |
treal_eq (treal_neg x1a) u & dest_real x1 x1a))" |
|
3857 |
by (import hollight DEF_real_neg) |
|
3858 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3859 |
definition real_add :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3860 |
"real_add == |
3861 |
%(x1::hollight.real) y1::hollight.real. |
|
3862 |
mk_real |
|
3863 |
(%u::hreal * hreal. |
|
3864 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3865 |
treal_eq (treal_add x1a y1a) u & |
|
3866 |
dest_real x1 x1a & dest_real y1 y1a)" |
|
3867 |
||
3868 |
lemma DEF_real_add: "real_add = |
|
3869 |
(%(x1::hollight.real) y1::hollight.real. |
|
3870 |
mk_real |
|
3871 |
(%u::hreal * hreal. |
|
3872 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3873 |
treal_eq (treal_add x1a y1a) u & |
|
3874 |
dest_real x1 x1a & dest_real y1 y1a))" |
|
3875 |
by (import hollight DEF_real_add) |
|
3876 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3877 |
definition real_mul :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3878 |
"real_mul == |
3879 |
%(x1::hollight.real) y1::hollight.real. |
|
3880 |
mk_real |
|
3881 |
(%u::hreal * hreal. |
|
3882 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3883 |
treal_eq (treal_mul x1a y1a) u & |
|
3884 |
dest_real x1 x1a & dest_real y1 y1a)" |
|
3885 |
||
3886 |
lemma DEF_real_mul: "real_mul = |
|
3887 |
(%(x1::hollight.real) y1::hollight.real. |
|
3888 |
mk_real |
|
3889 |
(%u::hreal * hreal. |
|
3890 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3891 |
treal_eq (treal_mul x1a y1a) u & |
|
3892 |
dest_real x1 x1a & dest_real y1 y1a))" |
|
3893 |
by (import hollight DEF_real_mul) |
|
3894 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3895 |
definition real_le :: "hollight.real => hollight.real => bool" where |
| 17644 | 3896 |
"real_le == |
3897 |
%(x1::hollight.real) y1::hollight.real. |
|
3898 |
SOME u::bool. |
|
3899 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3900 |
treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a" |
|
3901 |
||
3902 |
lemma DEF_real_le: "real_le = |
|
3903 |
(%(x1::hollight.real) y1::hollight.real. |
|
3904 |
SOME u::bool. |
|
3905 |
EX (x1a::hreal * hreal) y1a::hreal * hreal. |
|
3906 |
treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)" |
|
3907 |
by (import hollight DEF_real_le) |
|
3908 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3909 |
definition real_inv :: "hollight.real => hollight.real" where |
| 17644 | 3910 |
"real_inv == |
3911 |
%x::hollight.real. |
|
3912 |
mk_real |
|
3913 |
(%u::hreal * hreal. |
|
3914 |
EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa)" |
|
3915 |
||
3916 |
lemma DEF_real_inv: "real_inv = |
|
3917 |
(%x::hollight.real. |
|
3918 |
mk_real |
|
3919 |
(%u::hreal * hreal. |
|
3920 |
EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa))" |
|
3921 |
by (import hollight DEF_real_inv) |
|
3922 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3923 |
definition real_sub :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3924 |
"real_sub == %(u::hollight.real) ua::hollight.real. real_add u (real_neg ua)" |
3925 |
||
3926 |
lemma DEF_real_sub: "real_sub = (%(u::hollight.real) ua::hollight.real. real_add u (real_neg ua))" |
|
3927 |
by (import hollight DEF_real_sub) |
|
3928 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3929 |
definition real_lt :: "hollight.real => hollight.real => bool" where |
| 17644 | 3930 |
"real_lt == %(u::hollight.real) ua::hollight.real. ~ real_le ua u" |
3931 |
||
3932 |
lemma DEF_real_lt: "real_lt = (%(u::hollight.real) ua::hollight.real. ~ real_le ua u)" |
|
3933 |
by (import hollight DEF_real_lt) |
|
3934 |
||
3935 |
consts |
|
3936 |
real_ge :: "hollight.real => hollight.real => bool" |
|
3937 |
||
3938 |
defs |
|
3939 |
real_ge_def: "hollight.real_ge == %(u::hollight.real) ua::hollight.real. real_le ua u" |
|
3940 |
||
3941 |
lemma DEF_real_ge: "hollight.real_ge = (%(u::hollight.real) ua::hollight.real. real_le ua u)" |
|
3942 |
by (import hollight DEF_real_ge) |
|
3943 |
||
3944 |
consts |
|
3945 |
real_gt :: "hollight.real => hollight.real => bool" |
|
3946 |
||
3947 |
defs |
|
3948 |
real_gt_def: "hollight.real_gt == %(u::hollight.real) ua::hollight.real. real_lt ua u" |
|
3949 |
||
3950 |
lemma DEF_real_gt: "hollight.real_gt = (%(u::hollight.real) ua::hollight.real. real_lt ua u)" |
|
3951 |
by (import hollight DEF_real_gt) |
|
3952 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3953 |
definition real_abs :: "hollight.real => hollight.real" where |
| 17644 | 3954 |
"real_abs == |
| 17652 | 3955 |
%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u)" |
| 17644 | 3956 |
|
3957 |
lemma DEF_real_abs: "real_abs = |
|
| 17652 | 3958 |
(%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u))" |
| 17644 | 3959 |
by (import hollight DEF_real_abs) |
3960 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3961 |
definition real_pow :: "hollight.real => nat => hollight.real" where |
| 17644 | 3962 |
"real_pow == |
3963 |
SOME real_pow::hollight.real => nat => hollight.real. |
|
| 17652 | 3964 |
(ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & |
| 17644 | 3965 |
(ALL (x::hollight.real) n::nat. |
3966 |
real_pow x (Suc n) = real_mul x (real_pow x n))" |
|
3967 |
||
3968 |
lemma DEF_real_pow: "real_pow = |
|
3969 |
(SOME real_pow::hollight.real => nat => hollight.real. |
|
| 17652 | 3970 |
(ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & |
| 17644 | 3971 |
(ALL (x::hollight.real) n::nat. |
3972 |
real_pow x (Suc n) = real_mul x (real_pow x n)))" |
|
3973 |
by (import hollight DEF_real_pow) |
|
3974 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3975 |
definition real_div :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3976 |
"real_div == %(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua)" |
3977 |
||
3978 |
lemma DEF_real_div: "real_div = (%(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua))" |
|
3979 |
by (import hollight DEF_real_div) |
|
3980 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3981 |
definition real_max :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3982 |
"real_max == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u" |
3983 |
||
3984 |
lemma DEF_real_max: "real_max = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u)" |
|
3985 |
by (import hollight DEF_real_max) |
|
3986 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
3987 |
definition real_min :: "hollight.real => hollight.real => hollight.real" where |
| 17644 | 3988 |
"real_min == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua" |
3989 |
||
3990 |
lemma DEF_real_min: "real_min = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua)" |
|
3991 |
by (import hollight DEF_real_min) |
|
3992 |
||
3993 |
lemma REAL_HREAL_LEMMA1: "EX x::hreal => hollight.real. |
|
3994 |
(ALL xa::hollight.real. |
|
| 17652 | 3995 |
real_le (real_of_num 0) xa = (EX y::hreal. xa = x y)) & |
| 17644 | 3996 |
(ALL (y::hreal) z::hreal. hreal_le y z = real_le (x y) (x z))" |
3997 |
by (import hollight REAL_HREAL_LEMMA1) |
|
3998 |
||
3999 |
lemma REAL_HREAL_LEMMA2: "EX (x::hollight.real => hreal) r::hreal => hollight.real. |
|
4000 |
(ALL xa::hreal. x (r xa) = xa) & |
|
| 17652 | 4001 |
(ALL xa::hollight.real. real_le (real_of_num 0) xa --> r (x xa) = xa) & |
4002 |
(ALL x::hreal. real_le (real_of_num 0) (r x)) & |
|
| 17644 | 4003 |
(ALL (x::hreal) y::hreal. hreal_le x y = real_le (r x) (r y))" |
4004 |
by (import hollight REAL_HREAL_LEMMA2) |
|
4005 |
||
4006 |
lemma REAL_COMPLETE_SOMEPOS: "ALL P::hollight.real => bool. |
|
| 17652 | 4007 |
(EX x::hollight.real. P x & real_le (real_of_num 0) x) & |
| 17644 | 4008 |
(EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) --> |
4009 |
(EX M::hollight.real. |
|
4010 |
(ALL x::hollight.real. P x --> real_le x M) & |
|
4011 |
(ALL M'::hollight.real. |
|
4012 |
(ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))" |
|
4013 |
by (import hollight REAL_COMPLETE_SOMEPOS) |
|
4014 |
||
4015 |
lemma REAL_COMPLETE: "ALL P::hollight.real => bool. |
|
4016 |
Ex P & |
|
4017 |
(EX M::hollight.real. ALL x::hollight.real. P x --> real_le x M) --> |
|
4018 |
(EX M::hollight.real. |
|
4019 |
(ALL x::hollight.real. P x --> real_le x M) & |
|
4020 |
(ALL M'::hollight.real. |
|
4021 |
(ALL x::hollight.real. P x --> real_le x M') --> real_le M M'))" |
|
4022 |
by (import hollight REAL_COMPLETE) |
|
4023 |
||
4024 |
lemma REAL_ADD_AC: "real_add (m::hollight.real) (n::hollight.real) = real_add n m & |
|
4025 |
real_add (real_add m n) (p::hollight.real) = real_add m (real_add n p) & |
|
4026 |
real_add m (real_add n p) = real_add n (real_add m p)" |
|
4027 |
by (import hollight REAL_ADD_AC) |
|
4028 |
||
| 17652 | 4029 |
lemma REAL_ADD_RINV: "ALL x::hollight.real. real_add x (real_neg x) = real_of_num 0" |
| 17644 | 4030 |
by (import hollight REAL_ADD_RINV) |
4031 |
||
4032 |
lemma REAL_EQ_ADD_LCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4033 |
(real_add x y = real_add x z) = (y = z)" |
|
4034 |
by (import hollight REAL_EQ_ADD_LCANCEL) |
|
4035 |
||
4036 |
lemma REAL_EQ_ADD_RCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4037 |
(real_add x z = real_add y z) = (x = y)" |
|
4038 |
by (import hollight REAL_EQ_ADD_RCANCEL) |
|
4039 |
||
| 19093 | 4040 |
lemma REAL_MUL_RZERO: "ALL x::hollight.real. real_mul x (real_of_num 0) = real_of_num 0" |
4041 |
by (import hollight REAL_MUL_RZERO) |
|
4042 |
||
4043 |
lemma REAL_MUL_LZERO: "ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0" |
|
4044 |
by (import hollight REAL_MUL_LZERO) |
|
4045 |
||
| 17644 | 4046 |
lemma REAL_NEG_NEG: "ALL x::hollight.real. real_neg (real_neg x) = x" |
4047 |
by (import hollight REAL_NEG_NEG) |
|
4048 |
||
4049 |
lemma REAL_MUL_RNEG: "ALL (x::hollight.real) y::hollight.real. |
|
4050 |
real_mul x (real_neg y) = real_neg (real_mul x y)" |
|
4051 |
by (import hollight REAL_MUL_RNEG) |
|
4052 |
||
4053 |
lemma REAL_MUL_LNEG: "ALL (x::hollight.real) y::hollight.real. |
|
4054 |
real_mul (real_neg x) y = real_neg (real_mul x y)" |
|
4055 |
by (import hollight REAL_MUL_LNEG) |
|
4056 |
||
| 19093 | 4057 |
lemma REAL_NEG_ADD: "ALL (x::hollight.real) y::hollight.real. |
4058 |
real_neg (real_add x y) = real_add (real_neg x) (real_neg y)" |
|
4059 |
by (import hollight REAL_NEG_ADD) |
|
4060 |
||
| 17652 | 4061 |
lemma REAL_ADD_RID: "ALL x::hollight.real. real_add x (real_of_num 0) = x" |
| 17644 | 4062 |
by (import hollight REAL_ADD_RID) |
4063 |
||
| 19093 | 4064 |
lemma REAL_NEG_0: "real_neg (real_of_num 0) = real_of_num 0" |
4065 |
by (import hollight REAL_NEG_0) |
|
4066 |
||
| 17644 | 4067 |
lemma REAL_LE_LNEG: "ALL (x::hollight.real) y::hollight.real. |
| 17652 | 4068 |
real_le (real_neg x) y = real_le (real_of_num 0) (real_add x y)" |
| 17644 | 4069 |
by (import hollight REAL_LE_LNEG) |
4070 |
||
4071 |
lemma REAL_LE_NEG2: "ALL (x::hollight.real) y::hollight.real. |
|
4072 |
real_le (real_neg x) (real_neg y) = real_le y x" |
|
4073 |
by (import hollight REAL_LE_NEG2) |
|
4074 |
||
4075 |
lemma REAL_LE_RNEG: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4076 |
real_le x (real_neg y) = real_le (real_add x y) (real_of_num 0)" |
| 17644 | 4077 |
by (import hollight REAL_LE_RNEG) |
4078 |
||
4079 |
lemma REAL_OF_NUM_POW: "ALL (x::nat) n::nat. real_pow (real_of_num x) n = real_of_num (EXP x n)" |
|
4080 |
by (import hollight REAL_OF_NUM_POW) |
|
4081 |
||
4082 |
lemma REAL_POW_NEG: "ALL (x::hollight.real) n::nat. |
|
4083 |
real_pow (real_neg x) n = |
|
4084 |
COND (EVEN n) (real_pow x n) (real_neg (real_pow x n))" |
|
4085 |
by (import hollight REAL_POW_NEG) |
|
4086 |
||
4087 |
lemma REAL_ABS_NUM: "ALL x::nat. real_abs (real_of_num x) = real_of_num x" |
|
4088 |
by (import hollight REAL_ABS_NUM) |
|
4089 |
||
4090 |
lemma REAL_ABS_NEG: "ALL x::hollight.real. real_abs (real_neg x) = real_abs x" |
|
4091 |
by (import hollight REAL_ABS_NEG) |
|
4092 |
||
4093 |
lemma REAL_LTE_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_lt x xa | real_le xa x" |
|
4094 |
by (import hollight REAL_LTE_TOTAL) |
|
4095 |
||
4096 |
lemma REAL_LET_TOTAL: "ALL (x::hollight.real) xa::hollight.real. real_le x xa | real_lt xa x" |
|
4097 |
by (import hollight REAL_LET_TOTAL) |
|
4098 |
||
| 19093 | 4099 |
lemma REAL_LT_IMP_LE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> real_le x y" |
4100 |
by (import hollight REAL_LT_IMP_LE) |
|
4101 |
||
4102 |
lemma REAL_LTE_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4103 |
real_lt x y & real_le y z --> real_lt x z" |
|
4104 |
by (import hollight REAL_LTE_TRANS) |
|
4105 |
||
| 17644 | 4106 |
lemma REAL_LET_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4107 |
real_le x y & real_lt y z --> real_lt x z" |
|
4108 |
by (import hollight REAL_LET_TRANS) |
|
4109 |
||
4110 |
lemma REAL_LT_TRANS: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4111 |
real_lt x y & real_lt y z --> real_lt x z" |
|
4112 |
by (import hollight REAL_LT_TRANS) |
|
4113 |
||
4114 |
lemma REAL_LE_ADD: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4115 |
real_le (real_of_num 0) x & real_le (real_of_num 0) y --> |
4116 |
real_le (real_of_num 0) (real_add x y)" |
|
| 17644 | 4117 |
by (import hollight REAL_LE_ADD) |
4118 |
||
4119 |
lemma REAL_LTE_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_le y x)" |
|
4120 |
by (import hollight REAL_LTE_ANTISYM) |
|
4121 |
||
| 19093 | 4122 |
lemma REAL_SUB_LE: "ALL (x::hollight.real) xa::hollight.real. |
4123 |
real_le (real_of_num 0) (real_sub x xa) = real_le xa x" |
|
4124 |
by (import hollight REAL_SUB_LE) |
|
4125 |
||
4126 |
lemma REAL_NEG_SUB: "ALL (x::hollight.real) xa::hollight.real. |
|
4127 |
real_neg (real_sub x xa) = real_sub xa x" |
|
4128 |
by (import hollight REAL_NEG_SUB) |
|
4129 |
||
4130 |
lemma REAL_LE_LT: "ALL (x::hollight.real) xa::hollight.real. |
|
4131 |
real_le x xa = (real_lt x xa | x = xa)" |
|
4132 |
by (import hollight REAL_LE_LT) |
|
4133 |
||
4134 |
lemma REAL_SUB_LT: "ALL (x::hollight.real) xa::hollight.real. |
|
4135 |
real_lt (real_of_num 0) (real_sub x xa) = real_lt xa x" |
|
4136 |
by (import hollight REAL_SUB_LT) |
|
4137 |
||
4138 |
lemma REAL_NOT_LT: "ALL (x::hollight.real) xa::hollight.real. (~ real_lt x xa) = real_le xa x" |
|
4139 |
by (import hollight REAL_NOT_LT) |
|
4140 |
||
4141 |
lemma REAL_SUB_0: "ALL (x::hollight.real) y::hollight.real. |
|
4142 |
(real_sub x y = real_of_num 0) = (x = y)" |
|
4143 |
by (import hollight REAL_SUB_0) |
|
4144 |
||
4145 |
lemma REAL_LT_LE: "ALL (x::hollight.real) y::hollight.real. |
|
4146 |
real_lt x y = (real_le x y & x ~= y)" |
|
4147 |
by (import hollight REAL_LT_LE) |
|
4148 |
||
| 17644 | 4149 |
lemma REAL_LT_REFL: "ALL x::hollight.real. ~ real_lt x x" |
4150 |
by (import hollight REAL_LT_REFL) |
|
4151 |
||
| 19093 | 4152 |
lemma REAL_LTE_ADD: "ALL (x::hollight.real) y::hollight.real. |
4153 |
real_lt (real_of_num 0) x & real_le (real_of_num 0) y --> |
|
4154 |
real_lt (real_of_num 0) (real_add x y)" |
|
4155 |
by (import hollight REAL_LTE_ADD) |
|
4156 |
||
| 17644 | 4157 |
lemma REAL_LET_ADD: "ALL (x::hollight.real) y::hollight.real. |
| 17652 | 4158 |
real_le (real_of_num 0) x & real_lt (real_of_num 0) y --> |
4159 |
real_lt (real_of_num 0) (real_add x y)" |
|
| 17644 | 4160 |
by (import hollight REAL_LET_ADD) |
4161 |
||
| 19093 | 4162 |
lemma REAL_LT_ADD: "ALL (x::hollight.real) y::hollight.real. |
4163 |
real_lt (real_of_num 0) x & real_lt (real_of_num 0) y --> |
|
4164 |
real_lt (real_of_num 0) (real_add x y)" |
|
4165 |
by (import hollight REAL_LT_ADD) |
|
4166 |
||
| 17644 | 4167 |
lemma REAL_ENTIRE: "ALL (x::hollight.real) y::hollight.real. |
| 17652 | 4168 |
(real_mul x y = real_of_num 0) = (x = real_of_num 0 | y = real_of_num 0)" |
| 17644 | 4169 |
by (import hollight REAL_ENTIRE) |
4170 |
||
| 19093 | 4171 |
lemma REAL_LE_NEGTOTAL: "ALL x::hollight.real. |
4172 |
real_le (real_of_num 0) x | real_le (real_of_num 0) (real_neg x)" |
|
4173 |
by (import hollight REAL_LE_NEGTOTAL) |
|
4174 |
||
4175 |
lemma REAL_LE_SQUARE: "ALL x::hollight.real. real_le (real_of_num 0) (real_mul x x)" |
|
4176 |
by (import hollight REAL_LE_SQUARE) |
|
4177 |
||
4178 |
lemma REAL_MUL_RID: "ALL x::hollight.real. real_mul x (real_of_num (NUMERAL_BIT1 0)) = x" |
|
4179 |
by (import hollight REAL_MUL_RID) |
|
4180 |
||
| 17644 | 4181 |
lemma REAL_POW_2: "ALL x::hollight.real. |
| 17652 | 4182 |
real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = real_mul x x" |
| 17644 | 4183 |
by (import hollight REAL_POW_2) |
4184 |
||
4185 |
lemma REAL_POLY_CLAUSES: "(ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4186 |
real_add x (real_add y z) = real_add (real_add x y) z) & |
|
4187 |
(ALL (x::hollight.real) y::hollight.real. real_add x y = real_add y x) & |
|
| 17652 | 4188 |
(ALL x::hollight.real. real_add (real_of_num 0) x = x) & |
| 17644 | 4189 |
(ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4190 |
real_mul x (real_mul y z) = real_mul (real_mul x y) z) & |
|
4191 |
(ALL (x::hollight.real) y::hollight.real. real_mul x y = real_mul y x) & |
|
| 17652 | 4192 |
(ALL x::hollight.real. real_mul (real_of_num (NUMERAL_BIT1 0)) x = x) & |
4193 |
(ALL x::hollight.real. real_mul (real_of_num 0) x = real_of_num 0) & |
|
| 17644 | 4194 |
(ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
4195 |
real_mul x (real_add xa xb) = |
|
4196 |
real_add (real_mul x xa) (real_mul x xb)) & |
|
| 17652 | 4197 |
(ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & |
| 17644 | 4198 |
(ALL (x::hollight.real) xa::nat. |
4199 |
real_pow x (Suc xa) = real_mul x (real_pow x xa))" |
|
4200 |
by (import hollight REAL_POLY_CLAUSES) |
|
4201 |
||
4202 |
lemma REAL_POLY_NEG_CLAUSES: "(ALL x::hollight.real. |
|
| 17652 | 4203 |
real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x) & |
| 17644 | 4204 |
(ALL (x::hollight.real) xa::hollight.real. |
4205 |
real_sub x xa = |
|
| 17652 | 4206 |
real_add x (real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) xa))" |
| 17644 | 4207 |
by (import hollight REAL_POLY_NEG_CLAUSES) |
4208 |
||
| 19093 | 4209 |
lemma REAL_POS: "ALL x::nat. real_le (real_of_num 0) (real_of_num x)" |
4210 |
by (import hollight REAL_POS) |
|
4211 |
||
| 17644 | 4212 |
lemma REAL_OF_NUM_LT: "ALL (x::nat) xa::nat. real_lt (real_of_num x) (real_of_num xa) = < x xa" |
4213 |
by (import hollight REAL_OF_NUM_LT) |
|
4214 |
||
4215 |
lemma REAL_OF_NUM_GE: "ALL (x::nat) xa::nat. |
|
4216 |
hollight.real_ge (real_of_num x) (real_of_num xa) = >= x xa" |
|
4217 |
by (import hollight REAL_OF_NUM_GE) |
|
4218 |
||
4219 |
lemma REAL_OF_NUM_GT: "ALL (x::nat) xa::nat. |
|
4220 |
hollight.real_gt (real_of_num x) (real_of_num xa) = > x xa" |
|
4221 |
by (import hollight REAL_OF_NUM_GT) |
|
4222 |
||
4223 |
lemma REAL_OF_NUM_SUC: "ALL x::nat. |
|
| 17652 | 4224 |
real_add (real_of_num x) (real_of_num (NUMERAL_BIT1 0)) = |
| 17644 | 4225 |
real_of_num (Suc x)" |
4226 |
by (import hollight REAL_OF_NUM_SUC) |
|
4227 |
||
4228 |
lemma REAL_OF_NUM_SUB: "ALL (m::nat) n::nat. |
|
4229 |
<= m n --> real_sub (real_of_num n) (real_of_num m) = real_of_num (n - m)" |
|
4230 |
by (import hollight REAL_OF_NUM_SUB) |
|
4231 |
||
4232 |
lemma REAL_MUL_AC: "real_mul (m::hollight.real) (n::hollight.real) = real_mul n m & |
|
4233 |
real_mul (real_mul m n) (p::hollight.real) = real_mul m (real_mul n p) & |
|
4234 |
real_mul m (real_mul n p) = real_mul n (real_mul m p)" |
|
4235 |
by (import hollight REAL_MUL_AC) |
|
4236 |
||
4237 |
lemma REAL_ADD_RDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4238 |
real_mul (real_add x y) z = real_add (real_mul x z) (real_mul y z)" |
|
4239 |
by (import hollight REAL_ADD_RDISTRIB) |
|
4240 |
||
4241 |
lemma REAL_LT_LADD_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4242 |
real_lt y z --> real_lt (real_add x y) (real_add x z)" |
|
4243 |
by (import hollight REAL_LT_LADD_IMP) |
|
4244 |
||
4245 |
lemma REAL_LT_MUL: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4246 |
real_lt (real_of_num 0) x & real_lt (real_of_num 0) y --> |
4247 |
real_lt (real_of_num 0) (real_mul x y)" |
|
| 17644 | 4248 |
by (import hollight REAL_LT_MUL) |
4249 |
||
4250 |
lemma REAL_EQ_ADD_LCANCEL_0: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4251 |
(real_add x y = x) = (y = real_of_num 0)" |
| 17644 | 4252 |
by (import hollight REAL_EQ_ADD_LCANCEL_0) |
4253 |
||
4254 |
lemma REAL_EQ_ADD_RCANCEL_0: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4255 |
(real_add x y = y) = (x = real_of_num 0)" |
| 17644 | 4256 |
by (import hollight REAL_EQ_ADD_RCANCEL_0) |
4257 |
||
| 19093 | 4258 |
lemma REAL_LNEG_UNIQ: "ALL (x::hollight.real) y::hollight.real. |
4259 |
(real_add x y = real_of_num 0) = (x = real_neg y)" |
|
4260 |
by (import hollight REAL_LNEG_UNIQ) |
|
4261 |
||
4262 |
lemma REAL_RNEG_UNIQ: "ALL (x::hollight.real) y::hollight.real. |
|
4263 |
(real_add x y = real_of_num 0) = (y = real_neg x)" |
|
4264 |
by (import hollight REAL_RNEG_UNIQ) |
|
4265 |
||
4266 |
lemma REAL_NEG_LMUL: "ALL (x::hollight.real) y::hollight.real. |
|
4267 |
real_neg (real_mul x y) = real_mul (real_neg x) y" |
|
4268 |
by (import hollight REAL_NEG_LMUL) |
|
4269 |
||
4270 |
lemma REAL_NEG_RMUL: "ALL (x::hollight.real) y::hollight.real. |
|
4271 |
real_neg (real_mul x y) = real_mul x (real_neg y)" |
|
4272 |
by (import hollight REAL_NEG_RMUL) |
|
4273 |
||
4274 |
lemma REAL_NEGNEG: "ALL x::hollight.real. real_neg (real_neg x) = x" |
|
4275 |
by (import hollight REAL_NEGNEG) |
|
4276 |
||
4277 |
lemma REAL_NEG_MUL2: "ALL (x::hollight.real) y::hollight.real. |
|
4278 |
real_mul (real_neg x) (real_neg y) = real_mul x y" |
|
4279 |
by (import hollight REAL_NEG_MUL2) |
|
4280 |
||
4281 |
lemma REAL_LT_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4282 |
real_lt (real_add x y) (real_add x z) = real_lt y z" |
|
4283 |
by (import hollight REAL_LT_LADD) |
|
4284 |
||
4285 |
lemma REAL_LT_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4286 |
real_lt (real_add x z) (real_add y z) = real_lt x y" |
|
4287 |
by (import hollight REAL_LT_RADD) |
|
4288 |
||
4289 |
lemma REAL_LT_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_lt x y & real_lt y x)" |
|
4290 |
by (import hollight REAL_LT_ANTISYM) |
|
4291 |
||
4292 |
lemma REAL_LT_GT: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> ~ real_lt y x" |
|
4293 |
by (import hollight REAL_LT_GT) |
|
4294 |
||
| 17644 | 4295 |
lemma REAL_NOT_EQ: "ALL (x::hollight.real) y::hollight.real. |
4296 |
(x ~= y) = (real_lt x y | real_lt y x)" |
|
4297 |
by (import hollight REAL_NOT_EQ) |
|
4298 |
||
| 19093 | 4299 |
lemma REAL_LE_TOTAL: "ALL (x::hollight.real) y::hollight.real. real_le x y | real_le y x" |
4300 |
by (import hollight REAL_LE_TOTAL) |
|
4301 |
||
4302 |
lemma REAL_LE_REFL: "ALL x::hollight.real. real_le x x" |
|
4303 |
by (import hollight REAL_LE_REFL) |
|
4304 |
||
| 17644 | 4305 |
lemma REAL_LE_ANTISYM: "ALL (x::hollight.real) y::hollight.real. |
4306 |
(real_le x y & real_le y x) = (x = y)" |
|
4307 |
by (import hollight REAL_LE_ANTISYM) |
|
4308 |
||
4309 |
lemma REAL_LET_ANTISYM: "ALL (x::hollight.real) y::hollight.real. ~ (real_le x y & real_lt y x)" |
|
4310 |
by (import hollight REAL_LET_ANTISYM) |
|
4311 |
||
| 19093 | 4312 |
lemma REAL_NEG_LT0: "ALL x::hollight.real. |
4313 |
real_lt (real_neg x) (real_of_num 0) = real_lt (real_of_num 0) x" |
|
4314 |
by (import hollight REAL_NEG_LT0) |
|
4315 |
||
4316 |
lemma REAL_NEG_GT0: "ALL x::hollight.real. |
|
4317 |
real_lt (real_of_num 0) (real_neg x) = real_lt x (real_of_num 0)" |
|
4318 |
by (import hollight REAL_NEG_GT0) |
|
4319 |
||
4320 |
lemma REAL_NEG_LE0: "ALL x::hollight.real. |
|
4321 |
real_le (real_neg x) (real_of_num 0) = real_le (real_of_num 0) x" |
|
4322 |
by (import hollight REAL_NEG_LE0) |
|
4323 |
||
4324 |
lemma REAL_NEG_GE0: "ALL x::hollight.real. |
|
4325 |
real_le (real_of_num 0) (real_neg x) = real_le x (real_of_num 0)" |
|
4326 |
by (import hollight REAL_NEG_GE0) |
|
4327 |
||
| 17644 | 4328 |
lemma REAL_LT_TOTAL: "ALL (x::hollight.real) y::hollight.real. x = y | real_lt x y | real_lt y x" |
4329 |
by (import hollight REAL_LT_TOTAL) |
|
4330 |
||
| 19093 | 4331 |
lemma REAL_LT_NEGTOTAL: "ALL x::hollight.real. |
4332 |
x = real_of_num 0 | |
|
4333 |
real_lt (real_of_num 0) x | real_lt (real_of_num 0) (real_neg x)" |
|
4334 |
by (import hollight REAL_LT_NEGTOTAL) |
|
4335 |
||
| 17652 | 4336 |
lemma REAL_LE_01: "real_le (real_of_num 0) (real_of_num (NUMERAL_BIT1 0))" |
| 17644 | 4337 |
by (import hollight REAL_LE_01) |
4338 |
||
| 19093 | 4339 |
lemma REAL_LT_01: "real_lt (real_of_num 0) (real_of_num (NUMERAL_BIT1 0))" |
4340 |
by (import hollight REAL_LT_01) |
|
4341 |
||
4342 |
lemma REAL_LE_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4343 |
real_le (real_add x y) (real_add x z) = real_le y z" |
|
4344 |
by (import hollight REAL_LE_LADD) |
|
4345 |
||
4346 |
lemma REAL_LE_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4347 |
real_le (real_add x z) (real_add y z) = real_le x y" |
|
4348 |
by (import hollight REAL_LE_RADD) |
|
4349 |
||
4350 |
lemma REAL_LT_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
|
4351 |
z::hollight.real. |
|
4352 |
real_lt w x & real_lt y z --> real_lt (real_add w y) (real_add x z)" |
|
4353 |
by (import hollight REAL_LT_ADD2) |
|
4354 |
||
| 17644 | 4355 |
lemma REAL_LE_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
4356 |
z::hollight.real. |
|
4357 |
real_le w x & real_le y z --> real_le (real_add w y) (real_add x z)" |
|
4358 |
by (import hollight REAL_LE_ADD2) |
|
4359 |
||
4360 |
lemma REAL_LT_LNEG: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 4361 |
real_lt (real_neg x) xa = real_lt (real_of_num 0) (real_add x xa)" |
| 17644 | 4362 |
by (import hollight REAL_LT_LNEG) |
4363 |
||
4364 |
lemma REAL_LT_RNEG: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 4365 |
real_lt x (real_neg xa) = real_lt (real_add x xa) (real_of_num 0)" |
| 17644 | 4366 |
by (import hollight REAL_LT_RNEG) |
4367 |
||
| 19093 | 4368 |
lemma REAL_LT_ADDNEG: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4369 |
real_lt y (real_add x (real_neg z)) = real_lt (real_add y z) x" |
|
4370 |
by (import hollight REAL_LT_ADDNEG) |
|
4371 |
||
4372 |
lemma REAL_LT_ADDNEG2: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4373 |
real_lt (real_add x (real_neg y)) z = real_lt x (real_add z y)" |
|
4374 |
by (import hollight REAL_LT_ADDNEG2) |
|
4375 |
||
4376 |
lemma REAL_LT_ADD1: "ALL (x::hollight.real) y::hollight.real. |
|
4377 |
real_le x y --> real_lt x (real_add y (real_of_num (NUMERAL_BIT1 0)))" |
|
4378 |
by (import hollight REAL_LT_ADD1) |
|
4379 |
||
4380 |
lemma REAL_SUB_ADD: "ALL (x::hollight.real) y::hollight.real. real_add (real_sub x y) y = x" |
|
4381 |
by (import hollight REAL_SUB_ADD) |
|
4382 |
||
4383 |
lemma REAL_SUB_ADD2: "ALL (x::hollight.real) y::hollight.real. real_add y (real_sub x y) = x" |
|
4384 |
by (import hollight REAL_SUB_ADD2) |
|
4385 |
||
4386 |
lemma REAL_SUB_REFL: "ALL x::hollight.real. real_sub x x = real_of_num 0" |
|
4387 |
by (import hollight REAL_SUB_REFL) |
|
4388 |
||
4389 |
lemma REAL_LE_DOUBLE: "ALL x::hollight.real. |
|
4390 |
real_le (real_of_num 0) (real_add x x) = real_le (real_of_num 0) x" |
|
4391 |
by (import hollight REAL_LE_DOUBLE) |
|
4392 |
||
4393 |
lemma REAL_LE_NEGL: "ALL x::hollight.real. real_le (real_neg x) x = real_le (real_of_num 0) x" |
|
4394 |
by (import hollight REAL_LE_NEGL) |
|
4395 |
||
4396 |
lemma REAL_LE_NEGR: "ALL x::hollight.real. real_le x (real_neg x) = real_le x (real_of_num 0)" |
|
4397 |
by (import hollight REAL_LE_NEGR) |
|
4398 |
||
| 17652 | 4399 |
lemma REAL_NEG_EQ_0: "ALL x::hollight.real. (real_neg x = real_of_num 0) = (x = real_of_num 0)" |
| 17644 | 4400 |
by (import hollight REAL_NEG_EQ_0) |
4401 |
||
4402 |
lemma REAL_ADD_SUB: "ALL (x::hollight.real) y::hollight.real. real_sub (real_add x y) x = y" |
|
4403 |
by (import hollight REAL_ADD_SUB) |
|
4404 |
||
| 19093 | 4405 |
lemma REAL_NEG_EQ: "ALL (x::hollight.real) y::hollight.real. (real_neg x = y) = (x = real_neg y)" |
4406 |
by (import hollight REAL_NEG_EQ) |
|
4407 |
||
4408 |
lemma REAL_NEG_MINUS1: "ALL x::hollight.real. |
|
4409 |
real_neg x = real_mul (real_neg (real_of_num (NUMERAL_BIT1 0))) x" |
|
4410 |
by (import hollight REAL_NEG_MINUS1) |
|
4411 |
||
4412 |
lemma REAL_LT_IMP_NE: "ALL (x::hollight.real) y::hollight.real. real_lt x y --> x ~= y" |
|
4413 |
by (import hollight REAL_LT_IMP_NE) |
|
4414 |
||
| 17644 | 4415 |
lemma REAL_LE_ADDR: "ALL (x::hollight.real) y::hollight.real. |
| 17652 | 4416 |
real_le x (real_add x y) = real_le (real_of_num 0) y" |
| 17644 | 4417 |
by (import hollight REAL_LE_ADDR) |
4418 |
||
4419 |
lemma REAL_LE_ADDL: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4420 |
real_le y (real_add x y) = real_le (real_of_num 0) x" |
| 17644 | 4421 |
by (import hollight REAL_LE_ADDL) |
4422 |
||
4423 |
lemma REAL_LT_ADDR: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4424 |
real_lt x (real_add x y) = real_lt (real_of_num 0) y" |
| 17644 | 4425 |
by (import hollight REAL_LT_ADDR) |
4426 |
||
4427 |
lemma REAL_LT_ADDL: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4428 |
real_lt y (real_add x y) = real_lt (real_of_num 0) x" |
| 17644 | 4429 |
by (import hollight REAL_LT_ADDL) |
4430 |
||
| 19093 | 4431 |
lemma REAL_SUB_SUB: "ALL (x::hollight.real) y::hollight.real. |
4432 |
real_sub (real_sub x y) x = real_neg y" |
|
4433 |
by (import hollight REAL_SUB_SUB) |
|
4434 |
||
4435 |
lemma REAL_LT_ADD_SUB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4436 |
real_lt (real_add x y) z = real_lt x (real_sub z y)" |
|
4437 |
by (import hollight REAL_LT_ADD_SUB) |
|
4438 |
||
4439 |
lemma REAL_LT_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4440 |
real_lt (real_sub x y) z = real_lt x (real_add z y)" |
|
4441 |
by (import hollight REAL_LT_SUB_RADD) |
|
4442 |
||
4443 |
lemma REAL_LT_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4444 |
real_lt x (real_sub y z) = real_lt (real_add x z) y" |
|
4445 |
by (import hollight REAL_LT_SUB_LADD) |
|
4446 |
||
4447 |
lemma REAL_LE_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4448 |
real_le x (real_sub y z) = real_le (real_add x z) y" |
|
4449 |
by (import hollight REAL_LE_SUB_LADD) |
|
4450 |
||
4451 |
lemma REAL_LE_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4452 |
real_le (real_sub x y) z = real_le x (real_add z y)" |
|
4453 |
by (import hollight REAL_LE_SUB_RADD) |
|
4454 |
||
4455 |
lemma REAL_LT_NEG: "ALL (x::hollight.real) y::hollight.real. |
|
4456 |
real_lt (real_neg x) (real_neg y) = real_lt y x" |
|
4457 |
by (import hollight REAL_LT_NEG) |
|
4458 |
||
4459 |
lemma REAL_LE_NEG: "ALL (x::hollight.real) y::hollight.real. |
|
4460 |
real_le (real_neg x) (real_neg y) = real_le y x" |
|
4461 |
by (import hollight REAL_LE_NEG) |
|
4462 |
||
| 17644 | 4463 |
lemma REAL_ADD2_SUB2: "ALL (a::hollight.real) (b::hollight.real) (c::hollight.real) |
4464 |
d::hollight.real. |
|
4465 |
real_sub (real_add a b) (real_add c d) = |
|
4466 |
real_add (real_sub a c) (real_sub b d)" |
|
4467 |
by (import hollight REAL_ADD2_SUB2) |
|
4468 |
||
| 19093 | 4469 |
lemma REAL_SUB_LZERO: "ALL x::hollight.real. real_sub (real_of_num 0) x = real_neg x" |
4470 |
by (import hollight REAL_SUB_LZERO) |
|
4471 |
||
4472 |
lemma REAL_SUB_RZERO: "ALL x::hollight.real. real_sub x (real_of_num 0) = x" |
|
4473 |
by (import hollight REAL_SUB_RZERO) |
|
4474 |
||
| 17644 | 4475 |
lemma REAL_LET_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
4476 |
z::hollight.real. |
|
4477 |
real_le w x & real_lt y z --> real_lt (real_add w y) (real_add x z)" |
|
4478 |
by (import hollight REAL_LET_ADD2) |
|
4479 |
||
| 19093 | 4480 |
lemma REAL_LTE_ADD2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
4481 |
z::hollight.real. |
|
4482 |
real_lt w x & real_le y z --> real_lt (real_add w y) (real_add x z)" |
|
4483 |
by (import hollight REAL_LTE_ADD2) |
|
4484 |
||
4485 |
lemma REAL_SUB_LNEG: "ALL (x::hollight.real) y::hollight.real. |
|
4486 |
real_sub (real_neg x) y = real_neg (real_add x y)" |
|
4487 |
by (import hollight REAL_SUB_LNEG) |
|
4488 |
||
4489 |
lemma REAL_SUB_RNEG: "ALL (x::hollight.real) y::hollight.real. |
|
4490 |
real_sub x (real_neg y) = real_add x y" |
|
4491 |
by (import hollight REAL_SUB_RNEG) |
|
4492 |
||
4493 |
lemma REAL_SUB_NEG2: "ALL (x::hollight.real) y::hollight.real. |
|
4494 |
real_sub (real_neg x) (real_neg y) = real_sub y x" |
|
4495 |
by (import hollight REAL_SUB_NEG2) |
|
4496 |
||
4497 |
lemma REAL_SUB_TRIANGLE: "ALL (a::hollight.real) (b::hollight.real) c::hollight.real. |
|
4498 |
real_add (real_sub a b) (real_sub b c) = real_sub a c" |
|
4499 |
by (import hollight REAL_SUB_TRIANGLE) |
|
4500 |
||
| 17644 | 4501 |
lemma REAL_EQ_SUB_LADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4502 |
(x = real_sub y z) = (real_add x z = y)" |
|
4503 |
by (import hollight REAL_EQ_SUB_LADD) |
|
4504 |
||
4505 |
lemma REAL_EQ_SUB_RADD: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4506 |
(real_sub x y = z) = (x = real_add z y)" |
|
4507 |
by (import hollight REAL_EQ_SUB_RADD) |
|
4508 |
||
| 19093 | 4509 |
lemma REAL_SUB_SUB2: "ALL (x::hollight.real) y::hollight.real. real_sub x (real_sub x y) = y" |
4510 |
by (import hollight REAL_SUB_SUB2) |
|
4511 |
||
| 17644 | 4512 |
lemma REAL_ADD_SUB2: "ALL (x::hollight.real) y::hollight.real. |
4513 |
real_sub x (real_add x y) = real_neg y" |
|
4514 |
by (import hollight REAL_ADD_SUB2) |
|
4515 |
||
4516 |
lemma REAL_EQ_IMP_LE: "ALL (x::hollight.real) y::hollight.real. x = y --> real_le x y" |
|
4517 |
by (import hollight REAL_EQ_IMP_LE) |
|
4518 |
||
| 19093 | 4519 |
lemma REAL_POS_NZ: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0" |
4520 |
by (import hollight REAL_POS_NZ) |
|
4521 |
||
| 17644 | 4522 |
lemma REAL_DIFFSQ: "ALL (x::hollight.real) y::hollight.real. |
4523 |
real_mul (real_add x y) (real_sub x y) = |
|
4524 |
real_sub (real_mul x x) (real_mul y y)" |
|
4525 |
by (import hollight REAL_DIFFSQ) |
|
4526 |
||
4527 |
lemma REAL_EQ_NEG2: "ALL (x::hollight.real) y::hollight.real. (real_neg x = real_neg y) = (x = y)" |
|
4528 |
by (import hollight REAL_EQ_NEG2) |
|
4529 |
||
4530 |
lemma REAL_LT_NEG2: "ALL (x::hollight.real) y::hollight.real. |
|
4531 |
real_lt (real_neg x) (real_neg y) = real_lt y x" |
|
4532 |
by (import hollight REAL_LT_NEG2) |
|
4533 |
||
| 19093 | 4534 |
lemma REAL_SUB_LDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4535 |
real_mul x (real_sub y z) = real_sub (real_mul x y) (real_mul x z)" |
|
4536 |
by (import hollight REAL_SUB_LDISTRIB) |
|
4537 |
||
4538 |
lemma REAL_SUB_RDISTRIB: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4539 |
real_mul (real_sub x y) z = real_sub (real_mul x z) (real_mul y z)" |
|
4540 |
by (import hollight REAL_SUB_RDISTRIB) |
|
4541 |
||
| 17652 | 4542 |
lemma REAL_ABS_ZERO: "ALL x::hollight.real. (real_abs x = real_of_num 0) = (x = real_of_num 0)" |
| 17644 | 4543 |
by (import hollight REAL_ABS_ZERO) |
4544 |
||
| 17652 | 4545 |
lemma REAL_ABS_0: "real_abs (real_of_num 0) = real_of_num 0" |
| 17644 | 4546 |
by (import hollight REAL_ABS_0) |
4547 |
||
| 17652 | 4548 |
lemma REAL_ABS_1: "real_abs (real_of_num (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 0)" |
| 17644 | 4549 |
by (import hollight REAL_ABS_1) |
4550 |
||
4551 |
lemma REAL_ABS_TRIANGLE: "ALL (x::hollight.real) y::hollight.real. |
|
4552 |
real_le (real_abs (real_add x y)) (real_add (real_abs x) (real_abs y))" |
|
4553 |
by (import hollight REAL_ABS_TRIANGLE) |
|
4554 |
||
4555 |
lemma REAL_ABS_TRIANGLE_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4556 |
real_le (real_add (real_abs x) (real_abs (real_sub y x))) z --> |
|
4557 |
real_le (real_abs y) z" |
|
4558 |
by (import hollight REAL_ABS_TRIANGLE_LE) |
|
4559 |
||
4560 |
lemma REAL_ABS_TRIANGLE_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4561 |
real_lt (real_add (real_abs x) (real_abs (real_sub y x))) z --> |
|
4562 |
real_lt (real_abs y) z" |
|
4563 |
by (import hollight REAL_ABS_TRIANGLE_LT) |
|
4564 |
||
| 17652 | 4565 |
lemma REAL_ABS_POS: "ALL x::hollight.real. real_le (real_of_num 0) (real_abs x)" |
| 17644 | 4566 |
by (import hollight REAL_ABS_POS) |
4567 |
||
4568 |
lemma REAL_ABS_SUB: "ALL (x::hollight.real) y::hollight.real. |
|
4569 |
real_abs (real_sub x y) = real_abs (real_sub y x)" |
|
4570 |
by (import hollight REAL_ABS_SUB) |
|
4571 |
||
4572 |
lemma REAL_ABS_NZ: "ALL x::hollight.real. |
|
| 17652 | 4573 |
(x ~= real_of_num 0) = real_lt (real_of_num 0) (real_abs x)" |
| 17644 | 4574 |
by (import hollight REAL_ABS_NZ) |
4575 |
||
4576 |
lemma REAL_ABS_ABS: "ALL x::hollight.real. real_abs (real_abs x) = real_abs x" |
|
4577 |
by (import hollight REAL_ABS_ABS) |
|
4578 |
||
4579 |
lemma REAL_ABS_LE: "ALL x::hollight.real. real_le x (real_abs x)" |
|
4580 |
by (import hollight REAL_ABS_LE) |
|
4581 |
||
| 17652 | 4582 |
lemma REAL_ABS_REFL: "ALL x::hollight.real. (real_abs x = x) = real_le (real_of_num 0) x" |
| 17644 | 4583 |
by (import hollight REAL_ABS_REFL) |
4584 |
||
4585 |
lemma REAL_ABS_BETWEEN: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real. |
|
| 17652 | 4586 |
(real_lt (real_of_num 0) d & |
| 17644 | 4587 |
real_lt (real_sub x d) y & real_lt y (real_add x d)) = |
4588 |
real_lt (real_abs (real_sub y x)) d" |
|
4589 |
by (import hollight REAL_ABS_BETWEEN) |
|
4590 |
||
4591 |
lemma REAL_ABS_BOUND: "ALL (x::hollight.real) (y::hollight.real) d::hollight.real. |
|
4592 |
real_lt (real_abs (real_sub x y)) d --> real_lt y (real_add x d)" |
|
4593 |
by (import hollight REAL_ABS_BOUND) |
|
4594 |
||
4595 |
lemma REAL_ABS_STILLNZ: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4596 |
real_lt (real_abs (real_sub x y)) (real_abs y) --> x ~= real_of_num 0" |
| 17644 | 4597 |
by (import hollight REAL_ABS_STILLNZ) |
4598 |
||
4599 |
lemma REAL_ABS_CASES: "ALL x::hollight.real. |
|
| 17652 | 4600 |
x = real_of_num 0 | real_lt (real_of_num 0) (real_abs x)" |
| 17644 | 4601 |
by (import hollight REAL_ABS_CASES) |
4602 |
||
4603 |
lemma REAL_ABS_BETWEEN1: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4604 |
real_lt x z & real_lt (real_abs (real_sub y x)) (real_sub z x) --> |
|
4605 |
real_lt y z" |
|
4606 |
by (import hollight REAL_ABS_BETWEEN1) |
|
4607 |
||
4608 |
lemma REAL_ABS_SIGN: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4609 |
real_lt (real_abs (real_sub x y)) y --> real_lt (real_of_num 0) x" |
| 17644 | 4610 |
by (import hollight REAL_ABS_SIGN) |
4611 |
||
4612 |
lemma REAL_ABS_SIGN2: "ALL (x::hollight.real) y::hollight.real. |
|
4613 |
real_lt (real_abs (real_sub x y)) (real_neg y) --> |
|
| 17652 | 4614 |
real_lt x (real_of_num 0)" |
| 17644 | 4615 |
by (import hollight REAL_ABS_SIGN2) |
4616 |
||
4617 |
lemma REAL_ABS_CIRCLE: "ALL (x::hollight.real) (y::hollight.real) h::hollight.real. |
|
4618 |
real_lt (real_abs h) (real_sub (real_abs y) (real_abs x)) --> |
|
4619 |
real_lt (real_abs (real_add x h)) (real_abs y)" |
|
4620 |
by (import hollight REAL_ABS_CIRCLE) |
|
4621 |
||
| 19093 | 4622 |
lemma REAL_SUB_ABS: "ALL (x::hollight.real) y::hollight.real. |
4623 |
real_le (real_sub (real_abs x) (real_abs y)) (real_abs (real_sub x y))" |
|
4624 |
by (import hollight REAL_SUB_ABS) |
|
4625 |
||
| 17644 | 4626 |
lemma REAL_ABS_SUB_ABS: "ALL (x::hollight.real) y::hollight.real. |
4627 |
real_le (real_abs (real_sub (real_abs x) (real_abs y))) |
|
4628 |
(real_abs (real_sub x y))" |
|
4629 |
by (import hollight REAL_ABS_SUB_ABS) |
|
4630 |
||
4631 |
lemma REAL_ABS_BETWEEN2: "ALL (x0::hollight.real) (x::hollight.real) (y0::hollight.real) |
|
4632 |
y::hollight.real. |
|
4633 |
real_lt x0 y0 & |
|
4634 |
real_lt |
|
| 17652 | 4635 |
(real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) |
| 17644 | 4636 |
(real_abs (real_sub x x0))) |
4637 |
(real_sub y0 x0) & |
|
4638 |
real_lt |
|
| 17652 | 4639 |
(real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) |
| 17644 | 4640 |
(real_abs (real_sub y y0))) |
4641 |
(real_sub y0 x0) --> |
|
4642 |
real_lt x y" |
|
4643 |
by (import hollight REAL_ABS_BETWEEN2) |
|
4644 |
||
4645 |
lemma REAL_ABS_BOUNDS: "ALL (x::hollight.real) k::hollight.real. |
|
4646 |
real_le (real_abs x) k = (real_le (real_neg k) x & real_le x k)" |
|
4647 |
by (import hollight REAL_ABS_BOUNDS) |
|
4648 |
||
4649 |
lemma REAL_MIN_MAX: "ALL (x::hollight.real) y::hollight.real. |
|
4650 |
real_min x y = real_neg (real_max (real_neg x) (real_neg y))" |
|
4651 |
by (import hollight REAL_MIN_MAX) |
|
4652 |
||
4653 |
lemma REAL_MAX_MIN: "ALL (x::hollight.real) y::hollight.real. |
|
4654 |
real_max x y = real_neg (real_min (real_neg x) (real_neg y))" |
|
4655 |
by (import hollight REAL_MAX_MIN) |
|
4656 |
||
4657 |
lemma REAL_MAX_MAX: "ALL (x::hollight.real) y::hollight.real. |
|
4658 |
real_le x (real_max x y) & real_le y (real_max x y)" |
|
4659 |
by (import hollight REAL_MAX_MAX) |
|
4660 |
||
4661 |
lemma REAL_MIN_MIN: "ALL (x::hollight.real) y::hollight.real. |
|
4662 |
real_le (real_min x y) x & real_le (real_min x y) y" |
|
4663 |
by (import hollight REAL_MIN_MIN) |
|
4664 |
||
4665 |
lemma REAL_MAX_SYM: "ALL (x::hollight.real) y::hollight.real. real_max x y = real_max y x" |
|
4666 |
by (import hollight REAL_MAX_SYM) |
|
4667 |
||
4668 |
lemma REAL_MIN_SYM: "ALL (x::hollight.real) y::hollight.real. real_min x y = real_min y x" |
|
4669 |
by (import hollight REAL_MIN_SYM) |
|
4670 |
||
4671 |
lemma REAL_LE_MAX: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4672 |
real_le z (real_max x y) = (real_le z x | real_le z y)" |
|
4673 |
by (import hollight REAL_LE_MAX) |
|
4674 |
||
4675 |
lemma REAL_LE_MIN: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4676 |
real_le z (real_min x y) = (real_le z x & real_le z y)" |
|
4677 |
by (import hollight REAL_LE_MIN) |
|
4678 |
||
4679 |
lemma REAL_LT_MAX: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4680 |
real_lt z (real_max x y) = (real_lt z x | real_lt z y)" |
|
4681 |
by (import hollight REAL_LT_MAX) |
|
4682 |
||
4683 |
lemma REAL_LT_MIN: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4684 |
real_lt z (real_min x y) = (real_lt z x & real_lt z y)" |
|
4685 |
by (import hollight REAL_LT_MIN) |
|
4686 |
||
4687 |
lemma REAL_MAX_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4688 |
real_le (real_max x y) z = (real_le x z & real_le y z)" |
|
4689 |
by (import hollight REAL_MAX_LE) |
|
4690 |
||
4691 |
lemma REAL_MIN_LE: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4692 |
real_le (real_min x y) z = (real_le x z | real_le y z)" |
|
4693 |
by (import hollight REAL_MIN_LE) |
|
4694 |
||
4695 |
lemma REAL_MAX_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4696 |
real_lt (real_max x y) z = (real_lt x z & real_lt y z)" |
|
4697 |
by (import hollight REAL_MAX_LT) |
|
4698 |
||
4699 |
lemma REAL_MIN_LT: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4700 |
real_lt (real_min x y) z = (real_lt x z | real_lt y z)" |
|
4701 |
by (import hollight REAL_MIN_LT) |
|
4702 |
||
4703 |
lemma REAL_MAX_ASSOC: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4704 |
real_max x (real_max y z) = real_max (real_max x y) z" |
|
4705 |
by (import hollight REAL_MAX_ASSOC) |
|
4706 |
||
4707 |
lemma REAL_MIN_ASSOC: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
4708 |
real_min x (real_min y z) = real_min (real_min x y) z" |
|
4709 |
by (import hollight REAL_MIN_ASSOC) |
|
4710 |
||
4711 |
lemma REAL_MAX_ACI: "real_max (x::hollight.real) (y::hollight.real) = real_max y x & |
|
4712 |
real_max (real_max x y) (z::hollight.real) = real_max x (real_max y z) & |
|
4713 |
real_max x (real_max y z) = real_max y (real_max x z) & |
|
4714 |
real_max x x = x & real_max x (real_max x y) = real_max x y" |
|
4715 |
by (import hollight REAL_MAX_ACI) |
|
4716 |
||
4717 |
lemma REAL_MIN_ACI: "real_min (x::hollight.real) (y::hollight.real) = real_min y x & |
|
4718 |
real_min (real_min x y) (z::hollight.real) = real_min x (real_min y z) & |
|
4719 |
real_min x (real_min y z) = real_min y (real_min x z) & |
|
4720 |
real_min x x = x & real_min x (real_min x y) = real_min x y" |
|
4721 |
by (import hollight REAL_MIN_ACI) |
|
4722 |
||
4723 |
lemma REAL_ABS_MUL: "ALL (x::hollight.real) y::hollight.real. |
|
4724 |
real_abs (real_mul x y) = real_mul (real_abs x) (real_abs y)" |
|
4725 |
by (import hollight REAL_ABS_MUL) |
|
4726 |
||
4727 |
lemma REAL_POW_LE: "ALL (x::hollight.real) n::nat. |
|
| 17652 | 4728 |
real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_pow x n)" |
| 17644 | 4729 |
by (import hollight REAL_POW_LE) |
4730 |
||
4731 |
lemma REAL_POW_LT: "ALL (x::hollight.real) n::nat. |
|
| 17652 | 4732 |
real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_pow x n)" |
| 17644 | 4733 |
by (import hollight REAL_POW_LT) |
4734 |
||
4735 |
lemma REAL_ABS_POW: "ALL (x::hollight.real) n::nat. |
|
4736 |
real_abs (real_pow x n) = real_pow (real_abs x) n" |
|
4737 |
by (import hollight REAL_ABS_POW) |
|
4738 |
||
4739 |
lemma REAL_LE_LMUL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4740 |
real_le (real_of_num 0) x & real_le xa xb --> |
| 17644 | 4741 |
real_le (real_mul x xa) (real_mul x xb)" |
4742 |
by (import hollight REAL_LE_LMUL) |
|
4743 |
||
4744 |
lemma REAL_LE_RMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4745 |
real_le x y & real_le (real_of_num 0) z --> |
| 17644 | 4746 |
real_le (real_mul x z) (real_mul y z)" |
4747 |
by (import hollight REAL_LE_RMUL) |
|
4748 |
||
4749 |
lemma REAL_LT_LMUL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4750 |
real_lt (real_of_num 0) x & real_lt xa xb --> |
| 17644 | 4751 |
real_lt (real_mul x xa) (real_mul x xb)" |
4752 |
by (import hollight REAL_LT_LMUL) |
|
4753 |
||
4754 |
lemma REAL_LT_RMUL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4755 |
real_lt x y & real_lt (real_of_num 0) z --> |
| 17644 | 4756 |
real_lt (real_mul x z) (real_mul y z)" |
4757 |
by (import hollight REAL_LT_RMUL) |
|
4758 |
||
4759 |
lemma REAL_EQ_MUL_LCANCEL: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4760 |
(real_mul x y = real_mul x z) = (x = real_of_num 0 | y = z)" |
| 17644 | 4761 |
by (import hollight REAL_EQ_MUL_LCANCEL) |
4762 |
||
4763 |
lemma REAL_EQ_MUL_RCANCEL: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4764 |
(real_mul x xb = real_mul xa xb) = (x = xa | xb = real_of_num 0)" |
| 17644 | 4765 |
by (import hollight REAL_EQ_MUL_RCANCEL) |
4766 |
||
4767 |
lemma REAL_MUL_LINV_UNIQ: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4768 |
real_mul x y = real_of_num (NUMERAL_BIT1 0) --> real_inv y = x" |
| 17644 | 4769 |
by (import hollight REAL_MUL_LINV_UNIQ) |
4770 |
||
4771 |
lemma REAL_MUL_RINV_UNIQ: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 4772 |
real_mul x xa = real_of_num (NUMERAL_BIT1 0) --> real_inv x = xa" |
| 17644 | 4773 |
by (import hollight REAL_MUL_RINV_UNIQ) |
4774 |
||
4775 |
lemma REAL_INV_INV: "ALL x::hollight.real. real_inv (real_inv x) = x" |
|
4776 |
by (import hollight REAL_INV_INV) |
|
4777 |
||
| 17652 | 4778 |
lemma REAL_INV_EQ_0: "ALL x::hollight.real. (real_inv x = real_of_num 0) = (x = real_of_num 0)" |
| 17644 | 4779 |
by (import hollight REAL_INV_EQ_0) |
4780 |
||
4781 |
lemma REAL_LT_INV: "ALL x::hollight.real. |
|
| 17652 | 4782 |
real_lt (real_of_num 0) x --> real_lt (real_of_num 0) (real_inv x)" |
| 17644 | 4783 |
by (import hollight REAL_LT_INV) |
4784 |
||
4785 |
lemma REAL_LT_INV_EQ: "ALL x::hollight.real. |
|
| 17652 | 4786 |
real_lt (real_of_num 0) (real_inv x) = real_lt (real_of_num 0) x" |
| 17644 | 4787 |
by (import hollight REAL_LT_INV_EQ) |
4788 |
||
4789 |
lemma REAL_INV_NEG: "ALL x::hollight.real. real_inv (real_neg x) = real_neg (real_inv x)" |
|
4790 |
by (import hollight REAL_INV_NEG) |
|
4791 |
||
4792 |
lemma REAL_LE_INV_EQ: "ALL x::hollight.real. |
|
| 17652 | 4793 |
real_le (real_of_num 0) (real_inv x) = real_le (real_of_num 0) x" |
| 17644 | 4794 |
by (import hollight REAL_LE_INV_EQ) |
4795 |
||
4796 |
lemma REAL_LE_INV: "ALL x::hollight.real. |
|
| 17652 | 4797 |
real_le (real_of_num 0) x --> real_le (real_of_num 0) (real_inv x)" |
| 17644 | 4798 |
by (import hollight REAL_LE_INV) |
4799 |
||
| 19093 | 4800 |
lemma REAL_MUL_RINV: "ALL x::hollight.real. |
4801 |
x ~= real_of_num 0 --> |
|
4802 |
real_mul x (real_inv x) = real_of_num (NUMERAL_BIT1 0)" |
|
4803 |
by (import hollight REAL_MUL_RINV) |
|
4804 |
||
| 17652 | 4805 |
lemma REAL_INV_1: "real_inv (real_of_num (NUMERAL_BIT1 0)) = real_of_num (NUMERAL_BIT1 0)" |
| 17644 | 4806 |
by (import hollight REAL_INV_1) |
4807 |
||
| 17652 | 4808 |
lemma REAL_DIV_1: "ALL x::hollight.real. real_div x (real_of_num (NUMERAL_BIT1 0)) = x" |
| 17644 | 4809 |
by (import hollight REAL_DIV_1) |
4810 |
||
4811 |
lemma REAL_DIV_REFL: "ALL x::hollight.real. |
|
| 17652 | 4812 |
x ~= real_of_num 0 --> real_div x x = real_of_num (NUMERAL_BIT1 0)" |
| 17644 | 4813 |
by (import hollight REAL_DIV_REFL) |
4814 |
||
4815 |
lemma REAL_DIV_RMUL: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 4816 |
xa ~= real_of_num 0 --> real_mul (real_div x xa) xa = x" |
| 17644 | 4817 |
by (import hollight REAL_DIV_RMUL) |
4818 |
||
4819 |
lemma REAL_DIV_LMUL: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 4820 |
xa ~= real_of_num 0 --> real_mul xa (real_div x xa) = x" |
| 17644 | 4821 |
by (import hollight REAL_DIV_LMUL) |
4822 |
||
4823 |
lemma REAL_ABS_INV: "ALL x::hollight.real. real_abs (real_inv x) = real_inv (real_abs x)" |
|
4824 |
by (import hollight REAL_ABS_INV) |
|
4825 |
||
4826 |
lemma REAL_ABS_DIV: "ALL (x::hollight.real) xa::hollight.real. |
|
4827 |
real_abs (real_div x xa) = real_div (real_abs x) (real_abs xa)" |
|
4828 |
by (import hollight REAL_ABS_DIV) |
|
4829 |
||
4830 |
lemma REAL_INV_MUL: "ALL (x::hollight.real) y::hollight.real. |
|
4831 |
real_inv (real_mul x y) = real_mul (real_inv x) (real_inv y)" |
|
4832 |
by (import hollight REAL_INV_MUL) |
|
4833 |
||
4834 |
lemma REAL_INV_DIV: "ALL (x::hollight.real) xa::hollight.real. |
|
4835 |
real_inv (real_div x xa) = real_div xa x" |
|
4836 |
by (import hollight REAL_INV_DIV) |
|
4837 |
||
4838 |
lemma REAL_POW_MUL: "ALL (x::hollight.real) (y::hollight.real) n::nat. |
|
4839 |
real_pow (real_mul x y) n = real_mul (real_pow x n) (real_pow y n)" |
|
4840 |
by (import hollight REAL_POW_MUL) |
|
4841 |
||
4842 |
lemma REAL_POW_INV: "ALL (x::hollight.real) n::nat. |
|
4843 |
real_pow (real_inv x) n = real_inv (real_pow x n)" |
|
4844 |
by (import hollight REAL_POW_INV) |
|
4845 |
||
4846 |
lemma REAL_POW_DIV: "ALL (x::hollight.real) (xa::hollight.real) xb::nat. |
|
4847 |
real_pow (real_div x xa) xb = real_div (real_pow x xb) (real_pow xa xb)" |
|
4848 |
by (import hollight REAL_POW_DIV) |
|
4849 |
||
4850 |
lemma REAL_POW_ADD: "ALL (x::hollight.real) (m::nat) n::nat. |
|
4851 |
real_pow x (m + n) = real_mul (real_pow x m) (real_pow x n)" |
|
4852 |
by (import hollight REAL_POW_ADD) |
|
4853 |
||
4854 |
lemma REAL_POW_NZ: "ALL (x::hollight.real) n::nat. |
|
| 17652 | 4855 |
x ~= real_of_num 0 --> real_pow x n ~= real_of_num 0" |
| 17644 | 4856 |
by (import hollight REAL_POW_NZ) |
4857 |
||
4858 |
lemma REAL_POW_SUB: "ALL (x::hollight.real) (m::nat) n::nat. |
|
| 17652 | 4859 |
x ~= real_of_num 0 & <= m n --> |
| 17644 | 4860 |
real_pow x (n - m) = real_div (real_pow x n) (real_pow x m)" |
4861 |
by (import hollight REAL_POW_SUB) |
|
4862 |
||
| 17652 | 4863 |
lemma REAL_LT_IMP_NZ: "ALL x::hollight.real. real_lt (real_of_num 0) x --> x ~= real_of_num 0" |
| 17644 | 4864 |
by (import hollight REAL_LT_IMP_NZ) |
4865 |
||
4866 |
lemma REAL_LT_LCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4867 |
real_lt (real_of_num 0) x & real_lt (real_mul x y) (real_mul x z) --> |
| 17644 | 4868 |
real_lt y z" |
4869 |
by (import hollight REAL_LT_LCANCEL_IMP) |
|
4870 |
||
4871 |
lemma REAL_LT_RCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4872 |
real_lt (real_of_num 0) xb & real_lt (real_mul x xb) (real_mul xa xb) --> |
| 17644 | 4873 |
real_lt x xa" |
4874 |
by (import hollight REAL_LT_RCANCEL_IMP) |
|
4875 |
||
4876 |
lemma REAL_LE_LCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4877 |
real_lt (real_of_num 0) x & real_le (real_mul x y) (real_mul x z) --> |
| 17644 | 4878 |
real_le y z" |
4879 |
by (import hollight REAL_LE_LCANCEL_IMP) |
|
4880 |
||
4881 |
lemma REAL_LE_RCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4882 |
real_lt (real_of_num 0) xb & real_le (real_mul x xb) (real_mul xa xb) --> |
| 17644 | 4883 |
real_le x xa" |
4884 |
by (import hollight REAL_LE_RCANCEL_IMP) |
|
4885 |
||
| 19093 | 4886 |
lemma REAL_LE_RMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
4887 |
real_lt (real_of_num 0) z --> |
|
4888 |
real_le (real_mul x z) (real_mul y z) = real_le x y" |
|
4889 |
by (import hollight REAL_LE_RMUL_EQ) |
|
4890 |
||
| 17644 | 4891 |
lemma REAL_LE_LMUL_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
| 17652 | 4892 |
real_lt (real_of_num 0) z --> |
| 17644 | 4893 |
real_le (real_mul z x) (real_mul z y) = real_le x y" |
4894 |
by (import hollight REAL_LE_LMUL_EQ) |
|
4895 |
||
| 19093 | 4896 |
lemma REAL_LT_RMUL_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
4897 |
real_lt (real_of_num 0) xb --> |
|
4898 |
real_lt (real_mul x xb) (real_mul xa xb) = real_lt x xa" |
|
4899 |
by (import hollight REAL_LT_RMUL_EQ) |
|
4900 |
||
4901 |
lemma REAL_LT_LMUL_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
4902 |
real_lt (real_of_num 0) xb --> |
|
4903 |
real_lt (real_mul xb x) (real_mul xb xa) = real_lt x xa" |
|
4904 |
by (import hollight REAL_LT_LMUL_EQ) |
|
4905 |
||
| 17644 | 4906 |
lemma REAL_LE_RDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
| 17652 | 4907 |
real_lt (real_of_num 0) z --> |
| 17644 | 4908 |
real_le x (real_div y z) = real_le (real_mul x z) y" |
4909 |
by (import hollight REAL_LE_RDIV_EQ) |
|
4910 |
||
4911 |
lemma REAL_LE_LDIV_EQ: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 4912 |
real_lt (real_of_num 0) z --> |
| 17644 | 4913 |
real_le (real_div x z) y = real_le x (real_mul y z)" |
4914 |
by (import hollight REAL_LE_LDIV_EQ) |
|
4915 |
||
4916 |
lemma REAL_LT_RDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4917 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4918 |
real_lt x (real_div xa xb) = real_lt (real_mul x xb) xa" |
4919 |
by (import hollight REAL_LT_RDIV_EQ) |
|
4920 |
||
4921 |
lemma REAL_LT_LDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4922 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4923 |
real_lt (real_div x xb) xa = real_lt x (real_mul xa xb)" |
4924 |
by (import hollight REAL_LT_LDIV_EQ) |
|
4925 |
||
4926 |
lemma REAL_EQ_RDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4927 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4928 |
(x = real_div xa xb) = (real_mul x xb = xa)" |
4929 |
by (import hollight REAL_EQ_RDIV_EQ) |
|
4930 |
||
4931 |
lemma REAL_EQ_LDIV_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4932 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4933 |
(real_div x xb = xa) = (x = real_mul xa xb)" |
4934 |
by (import hollight REAL_EQ_LDIV_EQ) |
|
4935 |
||
4936 |
lemma REAL_LT_DIV2_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4937 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4938 |
real_lt (real_div x xb) (real_div xa xb) = real_lt x xa" |
4939 |
by (import hollight REAL_LT_DIV2_EQ) |
|
4940 |
||
4941 |
lemma REAL_LE_DIV2_EQ: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 4942 |
real_lt (real_of_num 0) xb --> |
| 17644 | 4943 |
real_le (real_div x xb) (real_div xa xb) = real_le x xa" |
4944 |
by (import hollight REAL_LE_DIV2_EQ) |
|
4945 |
||
4946 |
lemma REAL_MUL_2: "ALL x::hollight.real. |
|
| 17652 | 4947 |
real_mul (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x = real_add x x" |
| 17644 | 4948 |
by (import hollight REAL_MUL_2) |
4949 |
||
4950 |
lemma REAL_POW_EQ_0: "ALL (x::hollight.real) n::nat. |
|
| 17652 | 4951 |
(real_pow x n = real_of_num 0) = (x = real_of_num 0 & n ~= 0)" |
| 17644 | 4952 |
by (import hollight REAL_POW_EQ_0) |
4953 |
||
4954 |
lemma REAL_LE_MUL2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
|
4955 |
z::hollight.real. |
|
| 17652 | 4956 |
real_le (real_of_num 0) w & |
4957 |
real_le w x & real_le (real_of_num 0) y & real_le y z --> |
|
| 17644 | 4958 |
real_le (real_mul w y) (real_mul x z)" |
4959 |
by (import hollight REAL_LE_MUL2) |
|
4960 |
||
4961 |
lemma REAL_LT_MUL2: "ALL (w::hollight.real) (x::hollight.real) (y::hollight.real) |
|
4962 |
z::hollight.real. |
|
| 17652 | 4963 |
real_le (real_of_num 0) w & |
4964 |
real_lt w x & real_le (real_of_num 0) y & real_lt y z --> |
|
| 17644 | 4965 |
real_lt (real_mul w y) (real_mul x z)" |
4966 |
by (import hollight REAL_LT_MUL2) |
|
4967 |
||
4968 |
lemma REAL_LT_SQUARE: "ALL x::hollight.real. |
|
| 17652 | 4969 |
real_lt (real_of_num 0) (real_mul x x) = (x ~= real_of_num 0)" |
| 17644 | 4970 |
by (import hollight REAL_LT_SQUARE) |
4971 |
||
4972 |
lemma REAL_INV_LE_1: "ALL x::hollight.real. |
|
| 17652 | 4973 |
real_le (real_of_num (NUMERAL_BIT1 0)) x --> |
4974 |
real_le (real_inv x) (real_of_num (NUMERAL_BIT1 0))" |
|
| 17644 | 4975 |
by (import hollight REAL_INV_LE_1) |
4976 |
||
4977 |
lemma REAL_POW_LE_1: "ALL (n::nat) x::hollight.real. |
|
| 17652 | 4978 |
real_le (real_of_num (NUMERAL_BIT1 0)) x --> |
4979 |
real_le (real_of_num (NUMERAL_BIT1 0)) (real_pow x n)" |
|
| 17644 | 4980 |
by (import hollight REAL_POW_LE_1) |
4981 |
||
4982 |
lemma REAL_POW_1_LE: "ALL (n::nat) x::hollight.real. |
|
| 17652 | 4983 |
real_le (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) --> |
4984 |
real_le (real_pow x n) (real_of_num (NUMERAL_BIT1 0))" |
|
| 17644 | 4985 |
by (import hollight REAL_POW_1_LE) |
4986 |
||
| 17652 | 4987 |
lemma REAL_POW_1: "ALL x::hollight.real. real_pow x (NUMERAL_BIT1 0) = x" |
| 17644 | 4988 |
by (import hollight REAL_POW_1) |
4989 |
||
4990 |
lemma REAL_POW_ONE: "ALL n::nat. |
|
| 17652 | 4991 |
real_pow (real_of_num (NUMERAL_BIT1 0)) n = real_of_num (NUMERAL_BIT1 0)" |
| 17644 | 4992 |
by (import hollight REAL_POW_ONE) |
4993 |
||
4994 |
lemma REAL_LT_INV2: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 4995 |
real_lt (real_of_num 0) x & real_lt x y --> |
| 17644 | 4996 |
real_lt (real_inv y) (real_inv x)" |
4997 |
by (import hollight REAL_LT_INV2) |
|
4998 |
||
4999 |
lemma REAL_LE_INV2: "ALL (x::hollight.real) y::hollight.real. |
|
| 17652 | 5000 |
real_lt (real_of_num 0) x & real_le x y --> |
| 17644 | 5001 |
real_le (real_inv y) (real_inv x)" |
5002 |
by (import hollight REAL_LE_INV2) |
|
5003 |
||
5004 |
lemma REAL_INV_1_LE: "ALL x::hollight.real. |
|
| 17652 | 5005 |
real_lt (real_of_num 0) x & real_le x (real_of_num (NUMERAL_BIT1 0)) --> |
5006 |
real_le (real_of_num (NUMERAL_BIT1 0)) (real_inv x)" |
|
| 17644 | 5007 |
by (import hollight REAL_INV_1_LE) |
5008 |
||
5009 |
lemma REAL_SUB_INV: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 5010 |
x ~= real_of_num 0 & xa ~= real_of_num 0 --> |
| 17644 | 5011 |
real_sub (real_inv x) (real_inv xa) = |
5012 |
real_div (real_sub xa x) (real_mul x xa)" |
|
5013 |
by (import hollight REAL_SUB_INV) |
|
5014 |
||
5015 |
lemma REAL_DOWN: "ALL d::hollight.real. |
|
| 17652 | 5016 |
real_lt (real_of_num 0) d --> |
5017 |
(EX x::hollight.real. real_lt (real_of_num 0) x & real_lt x d)" |
|
| 17644 | 5018 |
by (import hollight REAL_DOWN) |
5019 |
||
5020 |
lemma REAL_DOWN2: "ALL (d1::hollight.real) d2::hollight.real. |
|
| 17652 | 5021 |
real_lt (real_of_num 0) d1 & real_lt (real_of_num 0) d2 --> |
| 17644 | 5022 |
(EX e::hollight.real. |
| 17652 | 5023 |
real_lt (real_of_num 0) e & real_lt e d1 & real_lt e d2)" |
| 17644 | 5024 |
by (import hollight REAL_DOWN2) |
5025 |
||
5026 |
lemma REAL_POW_LE2: "ALL (n::nat) (x::hollight.real) y::hollight.real. |
|
| 17652 | 5027 |
real_le (real_of_num 0) x & real_le x y --> |
| 17644 | 5028 |
real_le (real_pow x n) (real_pow y n)" |
5029 |
by (import hollight REAL_POW_LE2) |
|
5030 |
||
5031 |
lemma REAL_POW_MONO: "ALL (m::nat) (n::nat) x::hollight.real. |
|
| 17652 | 5032 |
real_le (real_of_num (NUMERAL_BIT1 0)) x & <= m n --> |
| 17644 | 5033 |
real_le (real_pow x m) (real_pow x n)" |
5034 |
by (import hollight REAL_POW_MONO) |
|
5035 |
||
5036 |
lemma REAL_POW_LT2: "ALL (n::nat) (x::hollight.real) y::hollight.real. |
|
| 17652 | 5037 |
n ~= 0 & real_le (real_of_num 0) x & real_lt x y --> |
| 17644 | 5038 |
real_lt (real_pow x n) (real_pow y n)" |
5039 |
by (import hollight REAL_POW_LT2) |
|
5040 |
||
5041 |
lemma REAL_POW_MONO_LT: "ALL (m::nat) (n::nat) x::hollight.real. |
|
| 17652 | 5042 |
real_lt (real_of_num (NUMERAL_BIT1 0)) x & < m n --> |
| 17644 | 5043 |
real_lt (real_pow x m) (real_pow x n)" |
5044 |
by (import hollight REAL_POW_MONO_LT) |
|
5045 |
||
5046 |
lemma REAL_POW_POW: "ALL (x::hollight.real) (m::nat) n::nat. |
|
5047 |
real_pow (real_pow x m) n = real_pow x (m * n)" |
|
5048 |
by (import hollight REAL_POW_POW) |
|
5049 |
||
5050 |
lemma REAL_EQ_RCANCEL_IMP: "ALL (x::hollight.real) (y::hollight.real) z::hollight.real. |
|
| 17652 | 5051 |
z ~= real_of_num 0 & real_mul x z = real_mul y z --> x = y" |
| 17644 | 5052 |
by (import hollight REAL_EQ_RCANCEL_IMP) |
5053 |
||
5054 |
lemma REAL_EQ_LCANCEL_IMP: "ALL (x::hollight.real) (xa::hollight.real) xb::hollight.real. |
|
| 17652 | 5055 |
xb ~= real_of_num 0 & real_mul xb x = real_mul xb xa --> x = xa" |
| 17644 | 5056 |
by (import hollight REAL_EQ_LCANCEL_IMP) |
5057 |
||
5058 |
lemma REAL_LT_DIV: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 5059 |
real_lt (real_of_num 0) x & real_lt (real_of_num 0) xa --> |
5060 |
real_lt (real_of_num 0) (real_div x xa)" |
|
| 17644 | 5061 |
by (import hollight REAL_LT_DIV) |
5062 |
||
5063 |
lemma REAL_LE_DIV: "ALL (x::hollight.real) xa::hollight.real. |
|
| 17652 | 5064 |
real_le (real_of_num 0) x & real_le (real_of_num 0) xa --> |
5065 |
real_le (real_of_num 0) (real_div x xa)" |
|
| 17644 | 5066 |
by (import hollight REAL_LE_DIV) |
5067 |
||
5068 |
lemma REAL_DIV_POW2: "ALL (x::hollight.real) (m::nat) n::nat. |
|
| 17652 | 5069 |
x ~= real_of_num 0 --> |
| 17644 | 5070 |
real_div (real_pow x m) (real_pow x n) = |
5071 |
COND (<= n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))" |
|
5072 |
by (import hollight REAL_DIV_POW2) |
|
5073 |
||
5074 |
lemma REAL_DIV_POW2_ALT: "ALL (x::hollight.real) (m::nat) n::nat. |
|
| 17652 | 5075 |
x ~= real_of_num 0 --> |
| 17644 | 5076 |
real_div (real_pow x m) (real_pow x n) = |
5077 |
COND (< n m) (real_pow x (m - n)) (real_inv (real_pow x (n - m)))" |
|
5078 |
by (import hollight REAL_DIV_POW2_ALT) |
|
5079 |
||
5080 |
lemma REAL_LT_POW2: "ALL x::nat. |
|
| 17652 | 5081 |
real_lt (real_of_num 0) |
5082 |
(real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) x)" |
|
| 17644 | 5083 |
by (import hollight REAL_LT_POW2) |
5084 |
||
5085 |
lemma REAL_LE_POW2: "ALL n::nat. |
|
| 17652 | 5086 |
real_le (real_of_num (NUMERAL_BIT1 0)) |
5087 |
(real_pow (real_of_num (NUMERAL_BIT0 (NUMERAL_BIT1 0))) n)" |
|
| 17644 | 5088 |
by (import hollight REAL_LE_POW2) |
5089 |
||
5090 |
lemma REAL_POW2_ABS: "ALL x::hollight.real. |
|
| 17652 | 5091 |
real_pow (real_abs x) (NUMERAL_BIT0 (NUMERAL_BIT1 0)) = |
5092 |
real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))" |
|
| 17644 | 5093 |
by (import hollight REAL_POW2_ABS) |
5094 |
||
5095 |
lemma REAL_LE_SQUARE_ABS: "ALL (x::hollight.real) y::hollight.real. |
|
5096 |
real_le (real_abs x) (real_abs y) = |
|
| 17652 | 5097 |
real_le (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))) |
5098 |
(real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0)))" |
|
| 17644 | 5099 |
by (import hollight REAL_LE_SQUARE_ABS) |
5100 |
||
| 19093 | 5101 |
lemma REAL_SOS_EQ_0: "ALL (x::hollight.real) y::hollight.real. |
5102 |
(real_add (real_pow x (NUMERAL_BIT0 (NUMERAL_BIT1 0))) |
|
5103 |
(real_pow y (NUMERAL_BIT0 (NUMERAL_BIT1 0))) = |
|
5104 |
real_of_num 0) = |
|
5105 |
(x = real_of_num 0 & y = real_of_num 0)" |
|
5106 |
by (import hollight REAL_SOS_EQ_0) |
|
5107 |
||
| 17644 | 5108 |
lemma REAL_WLOG_LE: "(ALL (x::hollight.real) y::hollight.real. |
5109 |
(P::hollight.real => hollight.real => bool) x y = P y x) & |
|
5110 |
(ALL (x::hollight.real) y::hollight.real. real_le x y --> P x y) --> |
|
5111 |
(ALL x::hollight.real. All (P x))" |
|
5112 |
by (import hollight REAL_WLOG_LE) |
|
5113 |
||
5114 |
lemma REAL_WLOG_LT: "(ALL x::hollight.real. (P::hollight.real => hollight.real => bool) x x) & |
|
5115 |
(ALL (x::hollight.real) y::hollight.real. P x y = P y x) & |
|
5116 |
(ALL (x::hollight.real) y::hollight.real. real_lt x y --> P x y) --> |
|
5117 |
(ALL x::hollight.real. All (P x))" |
|
5118 |
by (import hollight REAL_WLOG_LT) |
|
5119 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5120 |
definition mod_real :: "hollight.real => hollight.real => hollight.real => bool" where |
| 17644 | 5121 |
"mod_real == |
5122 |
%(u::hollight.real) (ua::hollight.real) ub::hollight.real. |
|
5123 |
EX q::hollight.real. real_sub ua ub = real_mul q u" |
|
5124 |
||
5125 |
lemma DEF_mod_real: "mod_real = |
|
5126 |
(%(u::hollight.real) (ua::hollight.real) ub::hollight.real. |
|
5127 |
EX q::hollight.real. real_sub ua ub = real_mul q u)" |
|
5128 |
by (import hollight DEF_mod_real) |
|
5129 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5130 |
definition DECIMAL :: "nat => nat => hollight.real" where |
| 17644 | 5131 |
"DECIMAL == %(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua)" |
5132 |
||
5133 |
lemma DEF_DECIMAL: "DECIMAL = (%(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua))" |
|
5134 |
by (import hollight DEF_DECIMAL) |
|
5135 |
||
| 17652 | 5136 |
lemma RAT_LEMMA1: "(y1::hollight.real) ~= real_of_num 0 & |
5137 |
(y2::hollight.real) ~= real_of_num 0 --> |
|
| 17644 | 5138 |
real_add (real_div (x1::hollight.real) y1) |
5139 |
(real_div (x2::hollight.real) y2) = |
|
5140 |
real_mul (real_add (real_mul x1 y2) (real_mul x2 y1)) |
|
5141 |
(real_mul (real_inv y1) (real_inv y2))" |
|
5142 |
by (import hollight RAT_LEMMA1) |
|
5143 |
||
| 17652 | 5144 |
lemma RAT_LEMMA2: "real_lt (real_of_num 0) (y1::hollight.real) & |
5145 |
real_lt (real_of_num 0) (y2::hollight.real) --> |
|
| 17644 | 5146 |
real_add (real_div (x1::hollight.real) y1) |
5147 |
(real_div (x2::hollight.real) y2) = |
|
5148 |
real_mul (real_add (real_mul x1 y2) (real_mul x2 y1)) |
|
5149 |
(real_mul (real_inv y1) (real_inv y2))" |
|
5150 |
by (import hollight RAT_LEMMA2) |
|
5151 |
||
| 17652 | 5152 |
lemma RAT_LEMMA3: "real_lt (real_of_num 0) (y1::hollight.real) & |
5153 |
real_lt (real_of_num 0) (y2::hollight.real) --> |
|
| 17644 | 5154 |
real_sub (real_div (x1::hollight.real) y1) |
5155 |
(real_div (x2::hollight.real) y2) = |
|
5156 |
real_mul (real_sub (real_mul x1 y2) (real_mul x2 y1)) |
|
5157 |
(real_mul (real_inv y1) (real_inv y2))" |
|
5158 |
by (import hollight RAT_LEMMA3) |
|
5159 |
||
| 17652 | 5160 |
lemma RAT_LEMMA4: "real_lt (real_of_num 0) (y1::hollight.real) & |
5161 |
real_lt (real_of_num 0) (y2::hollight.real) --> |
|
| 17644 | 5162 |
real_le (real_div (x1::hollight.real) y1) |
5163 |
(real_div (x2::hollight.real) y2) = |
|
5164 |
real_le (real_mul x1 y2) (real_mul x2 y1)" |
|
5165 |
by (import hollight RAT_LEMMA4) |
|
5166 |
||
| 17652 | 5167 |
lemma RAT_LEMMA5: "real_lt (real_of_num 0) (y1::hollight.real) & |
5168 |
real_lt (real_of_num 0) (y2::hollight.real) --> |
|
| 17644 | 5169 |
(real_div (x1::hollight.real) y1 = real_div (x2::hollight.real) y2) = |
5170 |
(real_mul x1 y2 = real_mul x2 y1)" |
|
5171 |
by (import hollight RAT_LEMMA5) |
|
5172 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5173 |
definition is_int :: "hollight.real => bool" where |
| 17644 | 5174 |
"is_int == |
5175 |
%u::hollight.real. |
|
5176 |
EX n::nat. u = real_of_num n | u = real_neg (real_of_num n)" |
|
5177 |
||
5178 |
lemma DEF_is_int: "is_int = |
|
5179 |
(%u::hollight.real. |
|
5180 |
EX n::nat. u = real_of_num n | u = real_neg (real_of_num n))" |
|
5181 |
by (import hollight DEF_is_int) |
|
5182 |
||
5183 |
typedef (open) int = "Collect is_int" morphisms "dest_int" "mk_int" |
|
| 17652 | 5184 |
apply (rule light_ex_imp_nonempty[where t="real_of_num (NUMERAL 0)"]) |
| 17644 | 5185 |
by (import hollight TYDEF_int) |
5186 |
||
5187 |
syntax |
|
5188 |
dest_int :: _ |
|
5189 |
||
5190 |
syntax |
|
5191 |
mk_int :: _ |
|
5192 |
||
5193 |
lemmas "TYDEF_int_@intern" = typedef_hol2hollight |
|
5194 |
[where a="a :: hollight.int" and r=r , |
|
5195 |
OF type_definition_int] |
|
5196 |
||
5197 |
lemma dest_int_rep: "ALL x::hollight.int. |
|
5198 |
EX n::nat. |
|
5199 |
dest_int x = real_of_num n | dest_int x = real_neg (real_of_num n)" |
|
5200 |
by (import hollight dest_int_rep) |
|
5201 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5202 |
definition int_le :: "hollight.int => hollight.int => bool" where |
| 17644 | 5203 |
"int_le == |
5204 |
%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua)" |
|
5205 |
||
5206 |
lemma DEF_int_le: "int_le = |
|
5207 |
(%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua))" |
|
5208 |
by (import hollight DEF_int_le) |
|
5209 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5210 |
definition int_lt :: "hollight.int => hollight.int => bool" where |
| 17644 | 5211 |
"int_lt == |
5212 |
%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua)" |
|
5213 |
||
5214 |
lemma DEF_int_lt: "int_lt = |
|
5215 |
(%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua))" |
|
5216 |
by (import hollight DEF_int_lt) |
|
5217 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5218 |
definition int_ge :: "hollight.int => hollight.int => bool" where |
| 17644 | 5219 |
"int_ge == |
5220 |
%(u::hollight.int) ua::hollight.int. |
|
5221 |
hollight.real_ge (dest_int u) (dest_int ua)" |
|
5222 |
||
5223 |
lemma DEF_int_ge: "int_ge = |
|
5224 |
(%(u::hollight.int) ua::hollight.int. |
|
5225 |
hollight.real_ge (dest_int u) (dest_int ua))" |
|
5226 |
by (import hollight DEF_int_ge) |
|
5227 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5228 |
definition int_gt :: "hollight.int => hollight.int => bool" where |
| 17644 | 5229 |
"int_gt == |
5230 |
%(u::hollight.int) ua::hollight.int. |
|
5231 |
hollight.real_gt (dest_int u) (dest_int ua)" |
|
5232 |
||
5233 |
lemma DEF_int_gt: "int_gt = |
|
5234 |
(%(u::hollight.int) ua::hollight.int. |
|
5235 |
hollight.real_gt (dest_int u) (dest_int ua))" |
|
5236 |
by (import hollight DEF_int_gt) |
|
5237 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5238 |
definition int_of_num :: "nat => hollight.int" where |
| 17644 | 5239 |
"int_of_num == %u::nat. mk_int (real_of_num u)" |
5240 |
||
5241 |
lemma DEF_int_of_num: "int_of_num = (%u::nat. mk_int (real_of_num u))" |
|
5242 |
by (import hollight DEF_int_of_num) |
|
5243 |
||
5244 |
lemma int_of_num_th: "ALL x::nat. dest_int (int_of_num x) = real_of_num x" |
|
5245 |
by (import hollight int_of_num_th) |
|
5246 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5247 |
definition int_neg :: "hollight.int => hollight.int" where |
| 17644 | 5248 |
"int_neg == %u::hollight.int. mk_int (real_neg (dest_int u))" |
5249 |
||
5250 |
lemma DEF_int_neg: "int_neg = (%u::hollight.int. mk_int (real_neg (dest_int u)))" |
|
5251 |
by (import hollight DEF_int_neg) |
|
5252 |
||
5253 |
lemma int_neg_th: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)" |
|
5254 |
by (import hollight int_neg_th) |
|
5255 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5256 |
definition int_add :: "hollight.int => hollight.int => hollight.int" where |
| 17644 | 5257 |
"int_add == |
5258 |
%(u::hollight.int) ua::hollight.int. |
|
5259 |
mk_int (real_add (dest_int u) (dest_int ua))" |
|
5260 |
||
5261 |
lemma DEF_int_add: "int_add = |
|
5262 |
(%(u::hollight.int) ua::hollight.int. |
|
5263 |
mk_int (real_add (dest_int u) (dest_int ua)))" |
|
5264 |
by (import hollight DEF_int_add) |
|
5265 |
||
5266 |
lemma int_add_th: "ALL (x::hollight.int) xa::hollight.int. |
|
5267 |
dest_int (int_add x xa) = real_add (dest_int x) (dest_int xa)" |
|
5268 |
by (import hollight int_add_th) |
|
5269 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5270 |
definition int_sub :: "hollight.int => hollight.int => hollight.int" where |
| 17644 | 5271 |
"int_sub == |
5272 |
%(u::hollight.int) ua::hollight.int. |
|
5273 |
mk_int (real_sub (dest_int u) (dest_int ua))" |
|
5274 |
||
5275 |
lemma DEF_int_sub: "int_sub = |
|
5276 |
(%(u::hollight.int) ua::hollight.int. |
|
5277 |
mk_int (real_sub (dest_int u) (dest_int ua)))" |
|
5278 |
by (import hollight DEF_int_sub) |
|
5279 |
||
5280 |
lemma int_sub_th: "ALL (x::hollight.int) xa::hollight.int. |
|
5281 |
dest_int (int_sub x xa) = real_sub (dest_int x) (dest_int xa)" |
|
5282 |
by (import hollight int_sub_th) |
|
5283 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5284 |
definition int_mul :: "hollight.int => hollight.int => hollight.int" where |
| 17644 | 5285 |
"int_mul == |
5286 |
%(u::hollight.int) ua::hollight.int. |
|
5287 |
mk_int (real_mul (dest_int u) (dest_int ua))" |
|
5288 |
||
5289 |
lemma DEF_int_mul: "int_mul = |
|
5290 |
(%(u::hollight.int) ua::hollight.int. |
|
5291 |
mk_int (real_mul (dest_int u) (dest_int ua)))" |
|
5292 |
by (import hollight DEF_int_mul) |
|
5293 |
||
5294 |
lemma int_mul_th: "ALL (x::hollight.int) y::hollight.int. |
|
5295 |
dest_int (int_mul x y) = real_mul (dest_int x) (dest_int y)" |
|
5296 |
by (import hollight int_mul_th) |
|
5297 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5298 |
definition int_abs :: "hollight.int => hollight.int" where |
| 17644 | 5299 |
"int_abs == %u::hollight.int. mk_int (real_abs (dest_int u))" |
5300 |
||
5301 |
lemma DEF_int_abs: "int_abs = (%u::hollight.int. mk_int (real_abs (dest_int u)))" |
|
5302 |
by (import hollight DEF_int_abs) |
|
5303 |
||
5304 |
lemma int_abs_th: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)" |
|
5305 |
by (import hollight int_abs_th) |
|
5306 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5307 |
definition int_max :: "hollight.int => hollight.int => hollight.int" where |
| 17644 | 5308 |
"int_max == |
5309 |
%(u::hollight.int) ua::hollight.int. |
|
5310 |
mk_int (real_max (dest_int u) (dest_int ua))" |
|
5311 |
||
5312 |
lemma DEF_int_max: "int_max = |
|
5313 |
(%(u::hollight.int) ua::hollight.int. |
|
5314 |
mk_int (real_max (dest_int u) (dest_int ua)))" |
|
5315 |
by (import hollight DEF_int_max) |
|
5316 |
||
5317 |
lemma int_max_th: "ALL (x::hollight.int) y::hollight.int. |
|
5318 |
dest_int (int_max x y) = real_max (dest_int x) (dest_int y)" |
|
5319 |
by (import hollight int_max_th) |
|
5320 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5321 |
definition int_min :: "hollight.int => hollight.int => hollight.int" where |
| 17644 | 5322 |
"int_min == |
5323 |
%(u::hollight.int) ua::hollight.int. |
|
5324 |
mk_int (real_min (dest_int u) (dest_int ua))" |
|
5325 |
||
5326 |
lemma DEF_int_min: "int_min = |
|
5327 |
(%(u::hollight.int) ua::hollight.int. |
|
5328 |
mk_int (real_min (dest_int u) (dest_int ua)))" |
|
5329 |
by (import hollight DEF_int_min) |
|
5330 |
||
5331 |
lemma int_min_th: "ALL (x::hollight.int) y::hollight.int. |
|
5332 |
dest_int (int_min x y) = real_min (dest_int x) (dest_int y)" |
|
5333 |
by (import hollight int_min_th) |
|
5334 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5335 |
definition int_pow :: "hollight.int => nat => hollight.int" where |
| 17644 | 5336 |
"int_pow == %(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua)" |
5337 |
||
5338 |
lemma DEF_int_pow: "int_pow = (%(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua))" |
|
5339 |
by (import hollight DEF_int_pow) |
|
5340 |
||
5341 |
lemma int_pow_th: "ALL (x::hollight.int) xa::nat. |
|
5342 |
dest_int (int_pow x xa) = real_pow (dest_int x) xa" |
|
5343 |
by (import hollight int_pow_th) |
|
5344 |
||
5345 |
lemma INT_IMAGE: "ALL x::hollight.int. |
|
5346 |
(EX n::nat. x = int_of_num n) | (EX n::nat. x = int_neg (int_of_num n))" |
|
5347 |
by (import hollight INT_IMAGE) |
|
5348 |
||
5349 |
lemma INT_LT_DISCRETE: "ALL (x::hollight.int) y::hollight.int. |
|
| 17652 | 5350 |
int_lt x y = int_le (int_add x (int_of_num (NUMERAL_BIT1 0))) y" |
| 17644 | 5351 |
by (import hollight INT_LT_DISCRETE) |
5352 |
||
5353 |
lemma INT_GT_DISCRETE: "ALL (x::hollight.int) xa::hollight.int. |
|
| 17652 | 5354 |
int_gt x xa = int_ge x (int_add xa (int_of_num (NUMERAL_BIT1 0)))" |
| 17644 | 5355 |
by (import hollight INT_GT_DISCRETE) |
5356 |
||
5357 |
lemma INT_FORALL_POS: "(ALL n::nat. (P::hollight.int => bool) (int_of_num n)) = |
|
| 17652 | 5358 |
(ALL i::hollight.int. int_le (int_of_num 0) i --> P i)" |
| 17644 | 5359 |
by (import hollight INT_FORALL_POS) |
5360 |
||
5361 |
lemma INT_ABS_MUL_1: "ALL (x::hollight.int) y::hollight.int. |
|
| 17652 | 5362 |
(int_abs (int_mul x y) = int_of_num (NUMERAL_BIT1 0)) = |
5363 |
(int_abs x = int_of_num (NUMERAL_BIT1 0) & |
|
5364 |
int_abs y = int_of_num (NUMERAL_BIT1 0))" |
|
| 17644 | 5365 |
by (import hollight INT_ABS_MUL_1) |
5366 |
||
| 17652 | 5367 |
lemma INT_POW: "int_pow (x::hollight.int) 0 = int_of_num (NUMERAL_BIT1 0) & |
| 17644 | 5368 |
(ALL xa::nat. int_pow x (Suc xa) = int_mul x (int_pow x xa))" |
5369 |
by (import hollight INT_POW) |
|
5370 |
||
5371 |
lemma INT_ABS: "ALL x::hollight.int. |
|
| 17652 | 5372 |
int_abs x = COND (int_le (int_of_num 0) x) x (int_neg x)" |
| 17644 | 5373 |
by (import hollight INT_ABS) |
5374 |
||
5375 |
lemma INT_GE: "ALL (x::hollight.int) xa::hollight.int. int_ge x xa = int_le xa x" |
|
5376 |
by (import hollight INT_GE) |
|
5377 |
||
5378 |
lemma INT_GT: "ALL (x::hollight.int) xa::hollight.int. int_gt x xa = int_lt xa x" |
|
5379 |
by (import hollight INT_GT) |
|
5380 |
||
5381 |
lemma INT_LT: "ALL (x::hollight.int) xa::hollight.int. int_lt x xa = (~ int_le xa x)" |
|
5382 |
by (import hollight INT_LT) |
|
5383 |
||
5384 |
lemma INT_ARCH: "ALL (x::hollight.int) d::hollight.int. |
|
| 17652 | 5385 |
d ~= int_of_num 0 --> (EX c::hollight.int. int_lt x (int_mul c d))" |
| 17644 | 5386 |
by (import hollight INT_ARCH) |
5387 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5388 |
definition mod_int :: "hollight.int => hollight.int => hollight.int => bool" where |
| 17644 | 5389 |
"mod_int == |
5390 |
%(u::hollight.int) (ua::hollight.int) ub::hollight.int. |
|
5391 |
EX q::hollight.int. int_sub ua ub = int_mul q u" |
|
5392 |
||
5393 |
lemma DEF_mod_int: "mod_int = |
|
5394 |
(%(u::hollight.int) (ua::hollight.int) ub::hollight.int. |
|
5395 |
EX q::hollight.int. int_sub ua ub = int_mul q u)" |
|
5396 |
by (import hollight DEF_mod_int) |
|
5397 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5398 |
definition IN :: "'A => ('A => bool) => bool" where
|
| 17644 | 5399 |
"IN == %(u::'A::type) ua::'A::type => bool. ua u" |
5400 |
||
5401 |
lemma DEF_IN: "IN = (%(u::'A::type) ua::'A::type => bool. ua u)" |
|
5402 |
by (import hollight DEF_IN) |
|
5403 |
||
5404 |
lemma EXTENSION: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5405 |
(x = xa) = (ALL xb::'A::type. IN xb x = IN xb xa)" |
|
5406 |
by (import hollight EXTENSION) |
|
5407 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5408 |
definition GSPEC :: "('A => bool) => 'A => bool" where
|
| 17644 | 5409 |
"GSPEC == %u::'A::type => bool. u" |
5410 |
||
5411 |
lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)" |
|
5412 |
by (import hollight DEF_GSPEC) |
|
5413 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5414 |
definition SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" where |
| 19093 | 5415 |
"SETSPEC == %(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub" |
5416 |
||
5417 |
lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub)" |
|
| 17644 | 5418 |
by (import hollight DEF_SETSPEC) |
5419 |
||
| 19093 | 5420 |
lemma IN_ELIM_THM: "(ALL (P::(bool => 'q_37089::type => bool) => bool) x::'q_37089::type. |
5421 |
IN x (GSPEC (%v::'q_37089::type. P (SETSPEC v))) = |
|
5422 |
P (%(p::bool) t::'q_37089::type. p & x = t)) & |
|
5423 |
(ALL (p::'q_37120::type => bool) x::'q_37120::type. |
|
| 17644 | 5424 |
IN x |
| 19093 | 5425 |
(GSPEC (%v::'q_37120::type. EX y::'q_37120::type. SETSPEC v (p y) y)) = |
| 17644 | 5426 |
p x) & |
| 19093 | 5427 |
(ALL (P::(bool => 'q_37148::type => bool) => bool) x::'q_37148::type. |
5428 |
GSPEC (%v::'q_37148::type. P (SETSPEC v)) x = |
|
5429 |
P (%(p::bool) t::'q_37148::type. p & x = t)) & |
|
5430 |
(ALL (p::'q_37177::type => bool) x::'q_37177::type. |
|
5431 |
GSPEC (%v::'q_37177::type. EX y::'q_37177::type. SETSPEC v (p y) y) x = |
|
| 17644 | 5432 |
p x) & |
| 19093 | 5433 |
(ALL (p::'q_37194::type => bool) x::'q_37194::type. IN x p = p x)" |
| 17644 | 5434 |
by (import hollight IN_ELIM_THM) |
5435 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5436 |
definition EMPTY :: "'A => bool" where |
| 17644 | 5437 |
"EMPTY == %x::'A::type. False" |
5438 |
||
5439 |
lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)" |
|
5440 |
by (import hollight DEF_EMPTY) |
|
5441 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5442 |
definition INSERT :: "'A => ('A => bool) => 'A => bool" where
|
| 17644 | 5443 |
"INSERT == %(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u" |
5444 |
||
5445 |
lemma DEF_INSERT: "INSERT = |
|
5446 |
(%(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u)" |
|
5447 |
by (import hollight DEF_INSERT) |
|
5448 |
||
5449 |
consts |
|
| 17652 | 5450 |
UNIV :: "'A => bool" |
| 17644 | 5451 |
|
5452 |
defs |
|
5453 |
UNIV_def: "hollight.UNIV == %x::'A::type. True" |
|
5454 |
||
5455 |
lemma DEF_UNIV: "hollight.UNIV = (%x::'A::type. True)" |
|
5456 |
by (import hollight DEF_UNIV) |
|
5457 |
||
5458 |
consts |
|
| 17652 | 5459 |
UNION :: "('A => bool) => ('A => bool) => 'A => bool"
|
| 17644 | 5460 |
|
5461 |
defs |
|
5462 |
UNION_def: "hollight.UNION == |
|
5463 |
%(u::'A::type => bool) ua::'A::type => bool. |
|
5464 |
GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x)" |
|
5465 |
||
5466 |
lemma DEF_UNION: "hollight.UNION = |
|
5467 |
(%(u::'A::type => bool) ua::'A::type => bool. |
|
5468 |
GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x))" |
|
5469 |
by (import hollight DEF_UNION) |
|
5470 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5471 |
definition UNIONS :: "(('A => bool) => bool) => 'A => bool" where
|
| 17644 | 5472 |
"UNIONS == |
5473 |
%u::('A::type => bool) => bool.
|
|
5474 |
GSPEC |
|
5475 |
(%ua::'A::type. |
|
5476 |
EX x::'A::type. |
|
5477 |
SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x)" |
|
5478 |
||
5479 |
lemma DEF_UNIONS: "UNIONS = |
|
5480 |
(%u::('A::type => bool) => bool.
|
|
5481 |
GSPEC |
|
5482 |
(%ua::'A::type. |
|
5483 |
EX x::'A::type. |
|
5484 |
SETSPEC ua (EX ua::'A::type => bool. IN ua u & IN x ua) x))" |
|
5485 |
by (import hollight DEF_UNIONS) |
|
5486 |
||
5487 |
consts |
|
| 17652 | 5488 |
INTER :: "('A => bool) => ('A => bool) => 'A => bool"
|
| 17644 | 5489 |
|
5490 |
defs |
|
5491 |
INTER_def: "hollight.INTER == |
|
5492 |
%(u::'A::type => bool) ua::'A::type => bool. |
|
5493 |
GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x)" |
|
5494 |
||
5495 |
lemma DEF_INTER: "hollight.INTER = |
|
5496 |
(%(u::'A::type => bool) ua::'A::type => bool. |
|
5497 |
GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x))" |
|
5498 |
by (import hollight DEF_INTER) |
|
5499 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5500 |
definition INTERS :: "(('A => bool) => bool) => 'A => bool" where
|
| 17644 | 5501 |
"INTERS == |
5502 |
%u::('A::type => bool) => bool.
|
|
5503 |
GSPEC |
|
5504 |
(%ua::'A::type. |
|
5505 |
EX x::'A::type. |
|
5506 |
SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x)" |
|
5507 |
||
5508 |
lemma DEF_INTERS: "INTERS = |
|
5509 |
(%u::('A::type => bool) => bool.
|
|
5510 |
GSPEC |
|
5511 |
(%ua::'A::type. |
|
5512 |
EX x::'A::type. |
|
5513 |
SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x))" |
|
5514 |
by (import hollight DEF_INTERS) |
|
5515 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5516 |
definition DIFF :: "('A => bool) => ('A => bool) => 'A => bool" where
|
| 17644 | 5517 |
"DIFF == |
5518 |
%(u::'A::type => bool) ua::'A::type => bool. |
|
5519 |
GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x)" |
|
5520 |
||
5521 |
lemma DEF_DIFF: "DIFF = |
|
5522 |
(%(u::'A::type => bool) ua::'A::type => bool. |
|
5523 |
GSPEC |
|
5524 |
(%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x))" |
|
5525 |
by (import hollight DEF_DIFF) |
|
5526 |
||
5527 |
lemma INSERT: "INSERT (x::'A::type) (s::'A::type => bool) = |
|
5528 |
GSPEC (%u::'A::type. EX y::'A::type. SETSPEC u (IN y s | y = x) y)" |
|
5529 |
by (import hollight INSERT) |
|
5530 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5531 |
definition DELETE :: "('A => bool) => 'A => 'A => bool" where
|
| 17644 | 5532 |
"DELETE == |
5533 |
%(u::'A::type => bool) ua::'A::type. |
|
5534 |
GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y)" |
|
5535 |
||
5536 |
lemma DEF_DELETE: "DELETE = |
|
5537 |
(%(u::'A::type => bool) ua::'A::type. |
|
5538 |
GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y))" |
|
5539 |
by (import hollight DEF_DELETE) |
|
5540 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5541 |
definition SUBSET :: "('A => bool) => ('A => bool) => bool" where
|
| 17644 | 5542 |
"SUBSET == |
5543 |
%(u::'A::type => bool) ua::'A::type => bool. |
|
5544 |
ALL x::'A::type. IN x u --> IN x ua" |
|
5545 |
||
5546 |
lemma DEF_SUBSET: "SUBSET = |
|
5547 |
(%(u::'A::type => bool) ua::'A::type => bool. |
|
5548 |
ALL x::'A::type. IN x u --> IN x ua)" |
|
5549 |
by (import hollight DEF_SUBSET) |
|
5550 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5551 |
definition PSUBSET :: "('A => bool) => ('A => bool) => bool" where
|
| 17644 | 5552 |
"PSUBSET == |
5553 |
%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua" |
|
5554 |
||
5555 |
lemma DEF_PSUBSET: "PSUBSET = |
|
5556 |
(%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua)" |
|
5557 |
by (import hollight DEF_PSUBSET) |
|
5558 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5559 |
definition DISJOINT :: "('A => bool) => ('A => bool) => bool" where
|
| 17644 | 5560 |
"DISJOINT == |
5561 |
%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY" |
|
5562 |
||
5563 |
lemma DEF_DISJOINT: "DISJOINT = |
|
5564 |
(%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY)" |
|
5565 |
by (import hollight DEF_DISJOINT) |
|
5566 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5567 |
definition SING :: "('A => bool) => bool" where
|
| 17644 | 5568 |
"SING == %u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY" |
5569 |
||
5570 |
lemma DEF_SING: "SING = (%u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY)" |
|
5571 |
by (import hollight DEF_SING) |
|
5572 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5573 |
definition FINITE :: "('A => bool) => bool" where
|
| 17644 | 5574 |
"FINITE == |
5575 |
%a::'A::type => bool. |
|
5576 |
ALL FINITE'::('A::type => bool) => bool.
|
|
5577 |
(ALL a::'A::type => bool. |
|
5578 |
a = EMPTY | |
|
5579 |
(EX (x::'A::type) s::'A::type => bool. |
|
5580 |
a = INSERT x s & FINITE' s) --> |
|
5581 |
FINITE' a) --> |
|
5582 |
FINITE' a" |
|
5583 |
||
5584 |
lemma DEF_FINITE: "FINITE = |
|
5585 |
(%a::'A::type => bool. |
|
5586 |
ALL FINITE'::('A::type => bool) => bool.
|
|
5587 |
(ALL a::'A::type => bool. |
|
5588 |
a = EMPTY | |
|
5589 |
(EX (x::'A::type) s::'A::type => bool. |
|
5590 |
a = INSERT x s & FINITE' s) --> |
|
5591 |
FINITE' a) --> |
|
5592 |
FINITE' a)" |
|
5593 |
by (import hollight DEF_FINITE) |
|
5594 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5595 |
definition INFINITE :: "('A => bool) => bool" where
|
| 17644 | 5596 |
"INFINITE == %u::'A::type => bool. ~ FINITE u" |
5597 |
||
5598 |
lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)" |
|
5599 |
by (import hollight DEF_INFINITE) |
|
5600 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5601 |
definition IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" where
|
| 17644 | 5602 |
"IMAGE == |
5603 |
%(u::'A::type => 'B::type) ua::'A::type => bool. |
|
5604 |
GSPEC |
|
5605 |
(%ub::'B::type. |
|
5606 |
EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y)" |
|
5607 |
||
5608 |
lemma DEF_IMAGE: "IMAGE = |
|
5609 |
(%(u::'A::type => 'B::type) ua::'A::type => bool. |
|
5610 |
GSPEC |
|
5611 |
(%ub::'B::type. |
|
5612 |
EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y))" |
|
5613 |
by (import hollight DEF_IMAGE) |
|
5614 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5615 |
definition INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
|
| 17644 | 5616 |
"INJ == |
5617 |
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5618 |
(ALL x::'A::type. IN x ua --> IN (u x) ub) & |
|
5619 |
(ALL (x::'A::type) y::'A::type. IN x ua & IN y ua & u x = u y --> x = y)" |
|
5620 |
||
5621 |
lemma DEF_INJ: "INJ = |
|
5622 |
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5623 |
(ALL x::'A::type. IN x ua --> IN (u x) ub) & |
|
5624 |
(ALL (x::'A::type) y::'A::type. |
|
5625 |
IN x ua & IN y ua & u x = u y --> x = y))" |
|
5626 |
by (import hollight DEF_INJ) |
|
5627 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5628 |
definition SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
|
| 17644 | 5629 |
"SURJ == |
5630 |
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5631 |
(ALL x::'A::type. IN x ua --> IN (u x) ub) & |
|
5632 |
(ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x))" |
|
5633 |
||
5634 |
lemma DEF_SURJ: "SURJ = |
|
5635 |
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5636 |
(ALL x::'A::type. IN x ua --> IN (u x) ub) & |
|
5637 |
(ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x)))" |
|
5638 |
by (import hollight DEF_SURJ) |
|
5639 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5640 |
definition BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where
|
| 17644 | 5641 |
"BIJ == |
5642 |
%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5643 |
INJ u ua ub & SURJ u ua ub" |
|
5644 |
||
5645 |
lemma DEF_BIJ: "BIJ = |
|
5646 |
(%(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. |
|
5647 |
INJ u ua ub & SURJ u ua ub)" |
|
5648 |
by (import hollight DEF_BIJ) |
|
5649 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5650 |
definition CHOICE :: "('A => bool) => 'A" where
|
| 17644 | 5651 |
"CHOICE == %u::'A::type => bool. SOME x::'A::type. IN x u" |
5652 |
||
5653 |
lemma DEF_CHOICE: "CHOICE = (%u::'A::type => bool. SOME x::'A::type. IN x u)" |
|
5654 |
by (import hollight DEF_CHOICE) |
|
5655 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5656 |
definition REST :: "('A => bool) => 'A => bool" where
|
| 17644 | 5657 |
"REST == %u::'A::type => bool. DELETE u (CHOICE u)" |
5658 |
||
5659 |
lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))" |
|
5660 |
by (import hollight DEF_REST) |
|
5661 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5662 |
definition CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" where
|
| 17644 | 5663 |
"CARD_GE == |
| 19093 | 5664 |
%(u::'q_37693::type => bool) ua::'q_37690::type => bool. |
5665 |
EX f::'q_37693::type => 'q_37690::type. |
|
5666 |
ALL y::'q_37690::type. |
|
5667 |
IN y ua --> (EX x::'q_37693::type. IN x u & y = f x)" |
|
| 17644 | 5668 |
|
5669 |
lemma DEF_CARD_GE: "CARD_GE = |
|
| 19093 | 5670 |
(%(u::'q_37693::type => bool) ua::'q_37690::type => bool. |
5671 |
EX f::'q_37693::type => 'q_37690::type. |
|
5672 |
ALL y::'q_37690::type. |
|
5673 |
IN y ua --> (EX x::'q_37693::type. IN x u & y = f x))" |
|
| 17644 | 5674 |
by (import hollight DEF_CARD_GE) |
5675 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5676 |
definition CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" where
|
| 17644 | 5677 |
"CARD_LE == |
| 19093 | 5678 |
%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u" |
| 17644 | 5679 |
|
5680 |
lemma DEF_CARD_LE: "CARD_LE = |
|
| 19093 | 5681 |
(%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u)" |
| 17644 | 5682 |
by (import hollight DEF_CARD_LE) |
5683 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5684 |
definition CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" where
|
| 17644 | 5685 |
"CARD_EQ == |
| 19093 | 5686 |
%(u::'q_37712::type => bool) ua::'q_37713::type => bool. |
| 17644 | 5687 |
CARD_LE u ua & CARD_LE ua u" |
5688 |
||
5689 |
lemma DEF_CARD_EQ: "CARD_EQ = |
|
| 19093 | 5690 |
(%(u::'q_37712::type => bool) ua::'q_37713::type => bool. |
| 17644 | 5691 |
CARD_LE u ua & CARD_LE ua u)" |
5692 |
by (import hollight DEF_CARD_EQ) |
|
5693 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5694 |
definition CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" where
|
| 17644 | 5695 |
"CARD_GT == |
| 19093 | 5696 |
%(u::'q_37727::type => bool) ua::'q_37728::type => bool. |
| 17644 | 5697 |
CARD_GE u ua & ~ CARD_GE ua u" |
5698 |
||
5699 |
lemma DEF_CARD_GT: "CARD_GT = |
|
| 19093 | 5700 |
(%(u::'q_37727::type => bool) ua::'q_37728::type => bool. |
| 17644 | 5701 |
CARD_GE u ua & ~ CARD_GE ua u)" |
5702 |
by (import hollight DEF_CARD_GT) |
|
5703 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5704 |
definition CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" where
|
| 17644 | 5705 |
"CARD_LT == |
| 19093 | 5706 |
%(u::'q_37743::type => bool) ua::'q_37744::type => bool. |
| 17644 | 5707 |
CARD_LE u ua & ~ CARD_LE ua u" |
5708 |
||
5709 |
lemma DEF_CARD_LT: "CARD_LT = |
|
| 19093 | 5710 |
(%(u::'q_37743::type => bool) ua::'q_37744::type => bool. |
| 17644 | 5711 |
CARD_LE u ua & ~ CARD_LE ua u)" |
5712 |
by (import hollight DEF_CARD_LT) |
|
5713 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
5714 |
definition COUNTABLE :: "('q_37757 => bool) => bool" where
|
| 19093 | 5715 |
"(op ==::(('q_37757::type => bool) => bool)
|
5716 |
=> (('q_37757::type => bool) => bool) => prop)
|
|
5717 |
(COUNTABLE::('q_37757::type => bool) => bool)
|
|
5718 |
((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool)
|
|
| 17644 | 5719 |
(hollight.UNIV::nat => bool))" |
5720 |
||
| 19093 | 5721 |
lemma DEF_COUNTABLE: "(op =::(('q_37757::type => bool) => bool)
|
5722 |
=> (('q_37757::type => bool) => bool) => bool)
|
|
5723 |
(COUNTABLE::('q_37757::type => bool) => bool)
|
|
5724 |
((CARD_GE::(nat => bool) => ('q_37757::type => bool) => bool)
|
|
| 17644 | 5725 |
(hollight.UNIV::nat => bool))" |
5726 |
by (import hollight DEF_COUNTABLE) |
|
5727 |
||
5728 |
lemma NOT_IN_EMPTY: "ALL x::'A::type. ~ IN x EMPTY" |
|
5729 |
by (import hollight NOT_IN_EMPTY) |
|
5730 |
||
5731 |
lemma IN_UNIV: "ALL x::'A::type. IN x hollight.UNIV" |
|
5732 |
by (import hollight IN_UNIV) |
|
5733 |
||
5734 |
lemma IN_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
5735 |
IN xb (hollight.UNION x xa) = (IN xb x | IN xb xa)" |
|
5736 |
by (import hollight IN_UNION) |
|
5737 |
||
5738 |
lemma IN_UNIONS: "ALL (x::('A::type => bool) => bool) xa::'A::type.
|
|
5739 |
IN xa (UNIONS x) = (EX t::'A::type => bool. IN t x & IN xa t)" |
|
5740 |
by (import hollight IN_UNIONS) |
|
5741 |
||
5742 |
lemma IN_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
5743 |
IN xb (hollight.INTER x xa) = (IN xb x & IN xb xa)" |
|
5744 |
by (import hollight IN_INTER) |
|
5745 |
||
5746 |
lemma IN_INTERS: "ALL (x::('A::type => bool) => bool) xa::'A::type.
|
|
5747 |
IN xa (INTERS x) = (ALL t::'A::type => bool. IN t x --> IN xa t)" |
|
5748 |
by (import hollight IN_INTERS) |
|
5749 |
||
5750 |
lemma IN_DIFF: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
5751 |
IN xb (DIFF x xa) = (IN xb x & ~ IN xb xa)" |
|
5752 |
by (import hollight IN_DIFF) |
|
5753 |
||
5754 |
lemma IN_INSERT: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. |
|
5755 |
IN x (INSERT xa xb) = (x = xa | IN x xb)" |
|
5756 |
by (import hollight IN_INSERT) |
|
5757 |
||
5758 |
lemma IN_DELETE: "ALL (x::'A::type => bool) (xa::'A::type) xb::'A::type. |
|
5759 |
IN xa (DELETE x xb) = (IN xa x & xa ~= xb)" |
|
5760 |
by (import hollight IN_DELETE) |
|
5761 |
||
5762 |
lemma IN_SING: "ALL (x::'A::type) xa::'A::type. IN x (INSERT xa EMPTY) = (x = xa)" |
|
5763 |
by (import hollight IN_SING) |
|
5764 |
||
5765 |
lemma IN_IMAGE: "ALL (x::'B::type) (xa::'A::type => bool) xb::'A::type => 'B::type. |
|
5766 |
IN x (IMAGE xb xa) = (EX xc::'A::type. x = xb xc & IN xc xa)" |
|
5767 |
by (import hollight IN_IMAGE) |
|
5768 |
||
5769 |
lemma IN_REST: "ALL (x::'A::type) xa::'A::type => bool. |
|
5770 |
IN x (REST xa) = (IN x xa & x ~= CHOICE xa)" |
|
5771 |
by (import hollight IN_REST) |
|
5772 |
||
5773 |
lemma CHOICE_DEF: "ALL x::'A::type => bool. x ~= EMPTY --> IN (CHOICE x) x" |
|
5774 |
by (import hollight CHOICE_DEF) |
|
5775 |
||
5776 |
lemma NOT_EQUAL_SETS: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5777 |
(x ~= xa) = (EX xb::'A::type. IN xb xa = (~ IN xb x))" |
|
5778 |
by (import hollight NOT_EQUAL_SETS) |
|
5779 |
||
5780 |
lemma MEMBER_NOT_EMPTY: "ALL x::'A::type => bool. (EX xa::'A::type. IN xa x) = (x ~= EMPTY)" |
|
5781 |
by (import hollight MEMBER_NOT_EMPTY) |
|
5782 |
||
5783 |
lemma UNIV_NOT_EMPTY: "(Not::bool => bool) |
|
5784 |
((op =::('A::type => bool) => ('A::type => bool) => bool)
|
|
5785 |
(hollight.UNIV::'A::type => bool) (EMPTY::'A::type => bool))" |
|
5786 |
by (import hollight UNIV_NOT_EMPTY) |
|
5787 |
||
5788 |
lemma EMPTY_NOT_UNIV: "(Not::bool => bool) |
|
5789 |
((op =::('A::type => bool) => ('A::type => bool) => bool)
|
|
5790 |
(EMPTY::'A::type => bool) (hollight.UNIV::'A::type => bool))" |
|
5791 |
by (import hollight EMPTY_NOT_UNIV) |
|
5792 |
||
5793 |
lemma EQ_UNIV: "(ALL x::'A::type. IN x (s::'A::type => bool)) = (s = hollight.UNIV)" |
|
5794 |
by (import hollight EQ_UNIV) |
|
5795 |
||
5796 |
lemma SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5797 |
SUBSET x xa & SUBSET xa xb --> SUBSET x xb" |
|
5798 |
by (import hollight SUBSET_TRANS) |
|
5799 |
||
5800 |
lemma SUBSET_REFL: "ALL x::'A::type => bool. SUBSET x x" |
|
5801 |
by (import hollight SUBSET_REFL) |
|
5802 |
||
5803 |
lemma SUBSET_ANTISYM: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5804 |
SUBSET x xa & SUBSET xa x --> x = xa" |
|
5805 |
by (import hollight SUBSET_ANTISYM) |
|
5806 |
||
5807 |
lemma EMPTY_SUBSET: "(All::(('A::type => bool) => bool) => bool)
|
|
5808 |
((SUBSET::('A::type => bool) => ('A::type => bool) => bool)
|
|
5809 |
(EMPTY::'A::type => bool))" |
|
5810 |
by (import hollight EMPTY_SUBSET) |
|
5811 |
||
5812 |
lemma SUBSET_EMPTY: "ALL x::'A::type => bool. SUBSET x EMPTY = (x = EMPTY)" |
|
5813 |
by (import hollight SUBSET_EMPTY) |
|
5814 |
||
5815 |
lemma SUBSET_UNIV: "ALL x::'A::type => bool. SUBSET x hollight.UNIV" |
|
5816 |
by (import hollight SUBSET_UNIV) |
|
5817 |
||
5818 |
lemma UNIV_SUBSET: "ALL x::'A::type => bool. SUBSET hollight.UNIV x = (x = hollight.UNIV)" |
|
5819 |
by (import hollight UNIV_SUBSET) |
|
5820 |
||
5821 |
lemma PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5822 |
PSUBSET x xa & PSUBSET xa xb --> PSUBSET x xb" |
|
5823 |
by (import hollight PSUBSET_TRANS) |
|
5824 |
||
5825 |
lemma PSUBSET_SUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5826 |
PSUBSET x xa & SUBSET xa xb --> PSUBSET x xb" |
|
5827 |
by (import hollight PSUBSET_SUBSET_TRANS) |
|
5828 |
||
5829 |
lemma SUBSET_PSUBSET_TRANS: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5830 |
SUBSET x xa & PSUBSET xa xb --> PSUBSET x xb" |
|
5831 |
by (import hollight SUBSET_PSUBSET_TRANS) |
|
5832 |
||
5833 |
lemma PSUBSET_IRREFL: "ALL x::'A::type => bool. ~ PSUBSET x x" |
|
5834 |
by (import hollight PSUBSET_IRREFL) |
|
5835 |
||
5836 |
lemma NOT_PSUBSET_EMPTY: "ALL x::'A::type => bool. ~ PSUBSET x EMPTY" |
|
5837 |
by (import hollight NOT_PSUBSET_EMPTY) |
|
5838 |
||
5839 |
lemma NOT_UNIV_PSUBSET: "ALL x::'A::type => bool. ~ PSUBSET hollight.UNIV x" |
|
5840 |
by (import hollight NOT_UNIV_PSUBSET) |
|
5841 |
||
5842 |
lemma PSUBSET_UNIV: "ALL x::'A::type => bool. |
|
5843 |
PSUBSET x hollight.UNIV = (EX xa::'A::type. ~ IN xa x)" |
|
5844 |
by (import hollight PSUBSET_UNIV) |
|
5845 |
||
5846 |
lemma UNION_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5847 |
hollight.UNION (hollight.UNION x xa) xb = |
|
5848 |
hollight.UNION x (hollight.UNION xa xb)" |
|
5849 |
by (import hollight UNION_ASSOC) |
|
5850 |
||
5851 |
lemma UNION_IDEMPOT: "ALL x::'A::type => bool. hollight.UNION x x = x" |
|
5852 |
by (import hollight UNION_IDEMPOT) |
|
5853 |
||
5854 |
lemma UNION_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5855 |
hollight.UNION x xa = hollight.UNION xa x" |
|
5856 |
by (import hollight UNION_COMM) |
|
5857 |
||
5858 |
lemma SUBSET_UNION: "(ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5859 |
SUBSET x (hollight.UNION x xa)) & |
|
5860 |
(ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5861 |
SUBSET x (hollight.UNION xa x))" |
|
5862 |
by (import hollight SUBSET_UNION) |
|
5863 |
||
5864 |
lemma SUBSET_UNION_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5865 |
SUBSET x xa = (hollight.UNION x xa = xa)" |
|
5866 |
by (import hollight SUBSET_UNION_ABSORPTION) |
|
5867 |
||
5868 |
lemma UNION_EMPTY: "(ALL x::'A::type => bool. hollight.UNION EMPTY x = x) & |
|
5869 |
(ALL x::'A::type => bool. hollight.UNION x EMPTY = x)" |
|
5870 |
by (import hollight UNION_EMPTY) |
|
5871 |
||
5872 |
lemma UNION_UNIV: "(ALL x::'A::type => bool. hollight.UNION hollight.UNIV x = hollight.UNIV) & |
|
5873 |
(ALL x::'A::type => bool. hollight.UNION x hollight.UNIV = hollight.UNIV)" |
|
5874 |
by (import hollight UNION_UNIV) |
|
5875 |
||
5876 |
lemma EMPTY_UNION: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5877 |
(hollight.UNION x xa = EMPTY) = (x = EMPTY & xa = EMPTY)" |
|
5878 |
by (import hollight EMPTY_UNION) |
|
5879 |
||
| 19093 | 5880 |
lemma UNION_SUBSET: "ALL (x::'q_38594::type => bool) (xa::'q_38594::type => bool) |
5881 |
xb::'q_38594::type => bool. |
|
| 17644 | 5882 |
SUBSET (hollight.UNION x xa) xb = (SUBSET x xb & SUBSET xa xb)" |
5883 |
by (import hollight UNION_SUBSET) |
|
5884 |
||
5885 |
lemma INTER_ASSOC: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5886 |
hollight.INTER (hollight.INTER x xa) xb = |
|
5887 |
hollight.INTER x (hollight.INTER xa xb)" |
|
5888 |
by (import hollight INTER_ASSOC) |
|
5889 |
||
5890 |
lemma INTER_IDEMPOT: "ALL x::'A::type => bool. hollight.INTER x x = x" |
|
5891 |
by (import hollight INTER_IDEMPOT) |
|
5892 |
||
5893 |
lemma INTER_COMM: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5894 |
hollight.INTER x xa = hollight.INTER xa x" |
|
5895 |
by (import hollight INTER_COMM) |
|
5896 |
||
5897 |
lemma INTER_SUBSET: "(ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5898 |
SUBSET (hollight.INTER x xa) x) & |
|
5899 |
(ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5900 |
SUBSET (hollight.INTER xa x) x)" |
|
5901 |
by (import hollight INTER_SUBSET) |
|
5902 |
||
5903 |
lemma SUBSET_INTER_ABSORPTION: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5904 |
SUBSET x xa = (hollight.INTER x xa = x)" |
|
5905 |
by (import hollight SUBSET_INTER_ABSORPTION) |
|
5906 |
||
5907 |
lemma INTER_EMPTY: "(ALL x::'A::type => bool. hollight.INTER EMPTY x = EMPTY) & |
|
5908 |
(ALL x::'A::type => bool. hollight.INTER x EMPTY = EMPTY)" |
|
5909 |
by (import hollight INTER_EMPTY) |
|
5910 |
||
5911 |
lemma INTER_UNIV: "(ALL x::'A::type => bool. hollight.INTER hollight.UNIV x = x) & |
|
5912 |
(ALL x::'A::type => bool. hollight.INTER x hollight.UNIV = x)" |
|
5913 |
by (import hollight INTER_UNIV) |
|
5914 |
||
5915 |
lemma UNION_OVER_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5916 |
hollight.INTER x (hollight.UNION xa xb) = |
|
5917 |
hollight.UNION (hollight.INTER x xa) (hollight.INTER x xb)" |
|
5918 |
by (import hollight UNION_OVER_INTER) |
|
5919 |
||
5920 |
lemma INTER_OVER_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5921 |
hollight.UNION x (hollight.INTER xa xb) = |
|
5922 |
hollight.INTER (hollight.UNION x xa) (hollight.UNION x xb)" |
|
5923 |
by (import hollight INTER_OVER_UNION) |
|
5924 |
||
5925 |
lemma IN_DISJOINT: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5926 |
DISJOINT x xa = (~ (EX xb::'A::type. IN xb x & IN xb xa))" |
|
5927 |
by (import hollight IN_DISJOINT) |
|
5928 |
||
5929 |
lemma DISJOINT_SYM: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5930 |
DISJOINT x xa = DISJOINT xa x" |
|
5931 |
by (import hollight DISJOINT_SYM) |
|
5932 |
||
5933 |
lemma DISJOINT_EMPTY: "ALL x::'A::type => bool. DISJOINT EMPTY x & DISJOINT x EMPTY" |
|
5934 |
by (import hollight DISJOINT_EMPTY) |
|
5935 |
||
5936 |
lemma DISJOINT_EMPTY_REFL: "ALL x::'A::type => bool. (x = EMPTY) = DISJOINT x x" |
|
5937 |
by (import hollight DISJOINT_EMPTY_REFL) |
|
5938 |
||
5939 |
lemma DISJOINT_UNION: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type => bool. |
|
5940 |
DISJOINT (hollight.UNION x xa) xb = (DISJOINT x xb & DISJOINT xa xb)" |
|
5941 |
by (import hollight DISJOINT_UNION) |
|
5942 |
||
5943 |
lemma DIFF_EMPTY: "ALL x::'A::type => bool. DIFF x EMPTY = x" |
|
5944 |
by (import hollight DIFF_EMPTY) |
|
5945 |
||
5946 |
lemma EMPTY_DIFF: "ALL x::'A::type => bool. DIFF EMPTY x = EMPTY" |
|
5947 |
by (import hollight EMPTY_DIFF) |
|
5948 |
||
5949 |
lemma DIFF_UNIV: "ALL x::'A::type => bool. DIFF x hollight.UNIV = EMPTY" |
|
5950 |
by (import hollight DIFF_UNIV) |
|
5951 |
||
5952 |
lemma DIFF_DIFF: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
5953 |
DIFF (DIFF x xa) xa = DIFF x xa" |
|
5954 |
by (import hollight DIFF_DIFF) |
|
5955 |
||
5956 |
lemma DIFF_EQ_EMPTY: "ALL x::'A::type => bool. DIFF x x = EMPTY" |
|
5957 |
by (import hollight DIFF_EQ_EMPTY) |
|
5958 |
||
| 19093 | 5959 |
lemma SUBSET_DIFF: "ALL (x::'q_39012::type => bool) xa::'q_39012::type => bool. |
| 17644 | 5960 |
SUBSET (DIFF x xa) x" |
5961 |
by (import hollight SUBSET_DIFF) |
|
5962 |
||
5963 |
lemma COMPONENT: "ALL (x::'A::type) s::'A::type => bool. IN x (INSERT x s)" |
|
5964 |
by (import hollight COMPONENT) |
|
5965 |
||
5966 |
lemma DECOMPOSITION: "ALL (s::'A::type => bool) x::'A::type. |
|
5967 |
IN x s = (EX t::'A::type => bool. s = INSERT x t & ~ IN x t)" |
|
5968 |
by (import hollight DECOMPOSITION) |
|
5969 |
||
5970 |
lemma SET_CASES: "ALL s::'A::type => bool. |
|
5971 |
s = EMPTY | |
|
5972 |
(EX (x::'A::type) t::'A::type => bool. s = INSERT x t & ~ IN x t)" |
|
5973 |
by (import hollight SET_CASES) |
|
5974 |
||
5975 |
lemma ABSORPTION: "ALL (x::'A::type) xa::'A::type => bool. IN x xa = (INSERT x xa = xa)" |
|
5976 |
by (import hollight ABSORPTION) |
|
5977 |
||
5978 |
lemma INSERT_INSERT: "ALL (x::'A::type) xa::'A::type => bool. INSERT x (INSERT x xa) = INSERT x xa" |
|
5979 |
by (import hollight INSERT_INSERT) |
|
5980 |
||
5981 |
lemma INSERT_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. |
|
5982 |
INSERT x (INSERT xa xb) = INSERT xa (INSERT x xb)" |
|
5983 |
by (import hollight INSERT_COMM) |
|
5984 |
||
5985 |
lemma INSERT_UNIV: "ALL x::'A::type. INSERT x hollight.UNIV = hollight.UNIV" |
|
5986 |
by (import hollight INSERT_UNIV) |
|
5987 |
||
5988 |
lemma NOT_INSERT_EMPTY: "ALL (x::'A::type) xa::'A::type => bool. INSERT x xa ~= EMPTY" |
|
5989 |
by (import hollight NOT_INSERT_EMPTY) |
|
5990 |
||
5991 |
lemma NOT_EMPTY_INSERT: "ALL (x::'A::type) xa::'A::type => bool. EMPTY ~= INSERT x xa" |
|
5992 |
by (import hollight NOT_EMPTY_INSERT) |
|
5993 |
||
5994 |
lemma INSERT_UNION: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool. |
|
5995 |
hollight.UNION (INSERT x s) t = |
|
5996 |
COND (IN x t) (hollight.UNION s t) (INSERT x (hollight.UNION s t))" |
|
5997 |
by (import hollight INSERT_UNION) |
|
5998 |
||
5999 |
lemma INSERT_UNION_EQ: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. |
|
6000 |
hollight.UNION (INSERT x xa) xb = INSERT x (hollight.UNION xa xb)" |
|
6001 |
by (import hollight INSERT_UNION_EQ) |
|
6002 |
||
6003 |
lemma INSERT_INTER: "ALL (x::'A::type) (s::'A::type => bool) t::'A::type => bool. |
|
6004 |
hollight.INTER (INSERT x s) t = |
|
6005 |
COND (IN x t) (INSERT x (hollight.INTER s t)) (hollight.INTER s t)" |
|
6006 |
by (import hollight INSERT_INTER) |
|
6007 |
||
6008 |
lemma DISJOINT_INSERT: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. |
|
6009 |
DISJOINT (INSERT x xa) xb = (DISJOINT xa xb & ~ IN x xb)" |
|
6010 |
by (import hollight DISJOINT_INSERT) |
|
6011 |
||
6012 |
lemma INSERT_SUBSET: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. |
|
6013 |
SUBSET (INSERT x xa) xb = (IN x xb & SUBSET xa xb)" |
|
6014 |
by (import hollight INSERT_SUBSET) |
|
6015 |
||
6016 |
lemma SUBSET_INSERT: "ALL (x::'A::type) xa::'A::type => bool. |
|
6017 |
~ IN x xa --> |
|
6018 |
(ALL xb::'A::type => bool. SUBSET xa (INSERT x xb) = SUBSET xa xb)" |
|
6019 |
by (import hollight SUBSET_INSERT) |
|
6020 |
||
6021 |
lemma INSERT_DIFF: "ALL (s::'A::type => bool) (t::'A::type => bool) x::'A::type. |
|
6022 |
DIFF (INSERT x s) t = COND (IN x t) (DIFF s t) (INSERT x (DIFF s t))" |
|
6023 |
by (import hollight INSERT_DIFF) |
|
6024 |
||
| 19093 | 6025 |
lemma INSERT_AC: "INSERT (x::'q_39468::type) |
6026 |
(INSERT (y::'q_39468::type) (s::'q_39468::type => bool)) = |
|
| 17644 | 6027 |
INSERT y (INSERT x s) & |
6028 |
INSERT x (INSERT x s) = INSERT x s" |
|
6029 |
by (import hollight INSERT_AC) |
|
6030 |
||
| 19093 | 6031 |
lemma INTER_ACI: "hollight.INTER (p::'q_39535::type => bool) (q::'q_39535::type => bool) = |
| 17644 | 6032 |
hollight.INTER q p & |
| 19093 | 6033 |
hollight.INTER (hollight.INTER p q) (r::'q_39535::type => bool) = |
| 17644 | 6034 |
hollight.INTER p (hollight.INTER q r) & |
6035 |
hollight.INTER p (hollight.INTER q r) = |
|
6036 |
hollight.INTER q (hollight.INTER p r) & |
|
6037 |
hollight.INTER p p = p & |
|
6038 |
hollight.INTER p (hollight.INTER p q) = hollight.INTER p q" |
|
6039 |
by (import hollight INTER_ACI) |
|
6040 |
||
| 19093 | 6041 |
lemma UNION_ACI: "hollight.UNION (p::'q_39601::type => bool) (q::'q_39601::type => bool) = |
| 17644 | 6042 |
hollight.UNION q p & |
| 19093 | 6043 |
hollight.UNION (hollight.UNION p q) (r::'q_39601::type => bool) = |
| 17644 | 6044 |
hollight.UNION p (hollight.UNION q r) & |
6045 |
hollight.UNION p (hollight.UNION q r) = |
|
6046 |
hollight.UNION q (hollight.UNION p r) & |
|
6047 |
hollight.UNION p p = p & |
|
6048 |
hollight.UNION p (hollight.UNION p q) = hollight.UNION p q" |
|
6049 |
by (import hollight UNION_ACI) |
|
6050 |
||
6051 |
lemma DELETE_NON_ELEMENT: "ALL (x::'A::type) xa::'A::type => bool. (~ IN x xa) = (DELETE xa x = xa)" |
|
6052 |
by (import hollight DELETE_NON_ELEMENT) |
|
6053 |
||
6054 |
lemma IN_DELETE_EQ: "ALL (s::'A::type => bool) (x::'A::type) x'::'A::type. |
|
6055 |
(IN x s = IN x' s) = (IN x (DELETE s x') = IN x' (DELETE s x))" |
|
6056 |
by (import hollight IN_DELETE_EQ) |
|
6057 |
||
6058 |
lemma EMPTY_DELETE: "ALL x::'A::type. DELETE EMPTY x = EMPTY" |
|
6059 |
by (import hollight EMPTY_DELETE) |
|
6060 |
||
6061 |
lemma DELETE_DELETE: "ALL (x::'A::type) xa::'A::type => bool. DELETE (DELETE xa x) x = DELETE xa x" |
|
6062 |
by (import hollight DELETE_DELETE) |
|
6063 |
||
6064 |
lemma DELETE_COMM: "ALL (x::'A::type) (xa::'A::type) xb::'A::type => bool. |
|
6065 |
DELETE (DELETE xb x) xa = DELETE (DELETE xb xa) x" |
|
6066 |
by (import hollight DELETE_COMM) |
|
6067 |
||
6068 |
lemma DELETE_SUBSET: "ALL (x::'A::type) xa::'A::type => bool. SUBSET (DELETE xa x) xa" |
|
6069 |
by (import hollight DELETE_SUBSET) |
|
6070 |
||
6071 |
lemma SUBSET_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. |
|
6072 |
SUBSET xa (DELETE xb x) = (~ IN x xa & SUBSET xa xb)" |
|
6073 |
by (import hollight SUBSET_DELETE) |
|
6074 |
||
6075 |
lemma SUBSET_INSERT_DELETE: "ALL (x::'A::type) (xa::'A::type => bool) xb::'A::type => bool. |
|
6076 |
SUBSET xa (INSERT x xb) = SUBSET (DELETE xa x) xb" |
|
6077 |
by (import hollight SUBSET_INSERT_DELETE) |
|
6078 |
||
6079 |
lemma DIFF_INSERT: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
6080 |
DIFF x (INSERT xb xa) = DIFF (DELETE x xb) xa" |
|
6081 |
by (import hollight DIFF_INSERT) |
|
6082 |
||
6083 |
lemma PSUBSET_INSERT_SUBSET: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
6084 |
PSUBSET x xa = (EX xb::'A::type. ~ IN xb x & SUBSET (INSERT xb x) xa)" |
|
6085 |
by (import hollight PSUBSET_INSERT_SUBSET) |
|
6086 |
||
6087 |
lemma PSUBSET_MEMBER: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
6088 |
PSUBSET x xa = (SUBSET x xa & (EX y::'A::type. IN y xa & ~ IN y x))" |
|
6089 |
by (import hollight PSUBSET_MEMBER) |
|
6090 |
||
6091 |
lemma DELETE_INSERT: "ALL (x::'A::type) (y::'A::type) s::'A::type => bool. |
|
6092 |
DELETE (INSERT x s) y = COND (x = y) (DELETE s y) (INSERT x (DELETE s y))" |
|
6093 |
by (import hollight DELETE_INSERT) |
|
6094 |
||
6095 |
lemma INSERT_DELETE: "ALL (x::'A::type) xa::'A::type => bool. |
|
6096 |
IN x xa --> INSERT x (DELETE xa x) = xa" |
|
6097 |
by (import hollight INSERT_DELETE) |
|
6098 |
||
6099 |
lemma DELETE_INTER: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
6100 |
hollight.INTER (DELETE x xb) xa = DELETE (hollight.INTER x xa) xb" |
|
6101 |
by (import hollight DELETE_INTER) |
|
6102 |
||
6103 |
lemma DISJOINT_DELETE_SYM: "ALL (x::'A::type => bool) (xa::'A::type => bool) xb::'A::type. |
|
6104 |
DISJOINT (DELETE x xb) xa = DISJOINT (DELETE xa xb) x" |
|
6105 |
by (import hollight DISJOINT_DELETE_SYM) |
|
6106 |
||
| 19093 | 6107 |
lemma UNIONS_0: "(op =::('q_40008::type => bool) => ('q_40008::type => bool) => bool)
|
6108 |
((UNIONS::(('q_40008::type => bool) => bool) => 'q_40008::type => bool)
|
|
6109 |
(EMPTY::('q_40008::type => bool) => bool))
|
|
6110 |
(EMPTY::'q_40008::type => bool)" |
|
| 17644 | 6111 |
by (import hollight UNIONS_0) |
6112 |
||
| 19093 | 6113 |
lemma UNIONS_1: "UNIONS (INSERT (s::'q_40014::type => bool) EMPTY) = s" |
| 17644 | 6114 |
by (import hollight UNIONS_1) |
6115 |
||
6116 |
lemma UNIONS_2: "UNIONS |
|
| 19093 | 6117 |
(INSERT (s::'q_40034::type => bool) |
6118 |
(INSERT (t::'q_40034::type => bool) EMPTY)) = |
|
| 17644 | 6119 |
hollight.UNION s t" |
6120 |
by (import hollight UNIONS_2) |
|
6121 |
||
6122 |
lemma UNIONS_INSERT: "UNIONS |
|
| 19093 | 6123 |
(INSERT (s::'q_40048::type => bool) |
6124 |
(u::('q_40048::type => bool) => bool)) =
|
|
| 17644 | 6125 |
hollight.UNION s (UNIONS u)" |
6126 |
by (import hollight UNIONS_INSERT) |
|
6127 |
||
| 19093 | 6128 |
lemma FORALL_IN_UNIONS: "ALL (x::'q_40090::type => bool) xa::('q_40090::type => bool) => bool.
|
6129 |
(ALL xb::'q_40090::type. IN xb (UNIONS xa) --> x xb) = |
|
6130 |
(ALL (t::'q_40090::type => bool) xb::'q_40090::type. |
|
| 17644 | 6131 |
IN t xa & IN xb t --> x xb)" |
6132 |
by (import hollight FORALL_IN_UNIONS) |
|
6133 |
||
| 19093 | 6134 |
lemma EMPTY_UNIONS: "ALL x::('q_40116::type => bool) => bool.
|
| 17644 | 6135 |
(UNIONS x = EMPTY) = |
| 19093 | 6136 |
(ALL xa::'q_40116::type => bool. IN xa x --> xa = EMPTY)" |
| 17644 | 6137 |
by (import hollight EMPTY_UNIONS) |
6138 |
||
| 19093 | 6139 |
lemma INTERS_0: "(op =::('q_40124::type => bool) => ('q_40124::type => bool) => bool)
|
6140 |
((INTERS::(('q_40124::type => bool) => bool) => 'q_40124::type => bool)
|
|
6141 |
(EMPTY::('q_40124::type => bool) => bool))
|
|
6142 |
(hollight.UNIV::'q_40124::type => bool)" |
|
6143 |
by (import hollight INTERS_0) |
|
6144 |
||
6145 |
lemma INTERS_1: "INTERS (INSERT (s::'q_40130::type => bool) EMPTY) = s" |
|
6146 |
by (import hollight INTERS_1) |
|
6147 |
||
6148 |
lemma INTERS_2: "INTERS |
|
6149 |
(INSERT (s::'q_40150::type => bool) |
|
6150 |
(INSERT (t::'q_40150::type => bool) EMPTY)) = |
|
6151 |
hollight.INTER s t" |
|
6152 |
by (import hollight INTERS_2) |
|
6153 |
||
6154 |
lemma INTERS_INSERT: "INTERS |
|
6155 |
(INSERT (s::'q_40164::type => bool) |
|
6156 |
(u::('q_40164::type => bool) => bool)) =
|
|
6157 |
hollight.INTER s (INTERS u)" |
|
6158 |
by (import hollight INTERS_INSERT) |
|
6159 |
||
6160 |
lemma IMAGE_CLAUSES: "IMAGE (f::'q_40190::type => 'q_40194::type) EMPTY = EMPTY & |
|
6161 |
IMAGE f (INSERT (x::'q_40190::type) (s::'q_40190::type => bool)) = |
|
| 17644 | 6162 |
INSERT (f x) (IMAGE f s)" |
6163 |
by (import hollight IMAGE_CLAUSES) |
|
6164 |
||
| 19093 | 6165 |
lemma IMAGE_UNION: "ALL (x::'q_40217::type => 'q_40228::type) (xa::'q_40217::type => bool) |
6166 |
xb::'q_40217::type => bool. |
|
| 17644 | 6167 |
IMAGE x (hollight.UNION xa xb) = hollight.UNION (IMAGE x xa) (IMAGE x xb)" |
6168 |
by (import hollight IMAGE_UNION) |
|
6169 |
||
| 19093 | 6170 |
lemma IMAGE_o: "ALL (x::'q_40261::type => 'q_40257::type) |
6171 |
(xa::'q_40252::type => 'q_40261::type) xb::'q_40252::type => bool. |
|
| 17644 | 6172 |
IMAGE (x o xa) xb = IMAGE x (IMAGE xa xb)" |
6173 |
by (import hollight IMAGE_o) |
|
6174 |
||
| 19093 | 6175 |
lemma IMAGE_SUBSET: "ALL (x::'q_40279::type => 'q_40290::type) (xa::'q_40279::type => bool) |
6176 |
xb::'q_40279::type => bool. |
|
| 17644 | 6177 |
SUBSET xa xb --> SUBSET (IMAGE x xa) (IMAGE x xb)" |
6178 |
by (import hollight IMAGE_SUBSET) |
|
6179 |
||
| 19093 | 6180 |
lemma IMAGE_DIFF_INJ: "(ALL (x::'q_40321::type) y::'q_40321::type. |
6181 |
(f::'q_40321::type => 'q_40332::type) x = f y --> x = y) --> |
|
6182 |
IMAGE f (DIFF (s::'q_40321::type => bool) (t::'q_40321::type => bool)) = |
|
| 17644 | 6183 |
DIFF (IMAGE f s) (IMAGE f t)" |
6184 |
by (import hollight IMAGE_DIFF_INJ) |
|
6185 |
||
| 19093 | 6186 |
lemma IMAGE_DELETE_INJ: "(ALL x::'q_40367::type. |
6187 |
(f::'q_40367::type => 'q_40366::type) x = f (a::'q_40367::type) --> |
|
| 17644 | 6188 |
x = a) --> |
| 19093 | 6189 |
IMAGE f (DELETE (s::'q_40367::type => bool) a) = DELETE (IMAGE f s) (f a)" |
| 17644 | 6190 |
by (import hollight IMAGE_DELETE_INJ) |
6191 |
||
| 19093 | 6192 |
lemma IMAGE_EQ_EMPTY: "ALL (x::'q_40390::type => 'q_40386::type) xa::'q_40390::type => bool. |
| 17644 | 6193 |
(IMAGE x xa = EMPTY) = (xa = EMPTY)" |
6194 |
by (import hollight IMAGE_EQ_EMPTY) |
|
6195 |
||
| 19093 | 6196 |
lemma FORALL_IN_IMAGE: "ALL (x::'q_40426::type => 'q_40425::type) xa::'q_40426::type => bool. |
6197 |
(ALL xb::'q_40425::type. |
|
6198 |
IN xb (IMAGE x xa) --> (P::'q_40425::type => bool) xb) = |
|
6199 |
(ALL xb::'q_40426::type. IN xb xa --> P (x xb))" |
|
| 17644 | 6200 |
by (import hollight FORALL_IN_IMAGE) |
6201 |
||
| 19093 | 6202 |
lemma EXISTS_IN_IMAGE: "ALL (x::'q_40462::type => 'q_40461::type) xa::'q_40462::type => bool. |
6203 |
(EX xb::'q_40461::type. |
|
6204 |
IN xb (IMAGE x xa) & (P::'q_40461::type => bool) xb) = |
|
6205 |
(EX xb::'q_40462::type. IN xb xa & P (x xb))" |
|
| 17644 | 6206 |
by (import hollight EXISTS_IN_IMAGE) |
6207 |
||
6208 |
lemma SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'B::type => bool) t::'A::type => bool. |
|
6209 |
SUBSET s (IMAGE f t) = |
|
6210 |
(EX x::'A::type => bool. SUBSET x t & s = IMAGE f x)" |
|
6211 |
by (import hollight SUBSET_IMAGE) |
|
6212 |
||
| 19093 | 6213 |
lemma IMAGE_CONST: "ALL (s::'q_40548::type => bool) c::'q_40553::type. |
6214 |
IMAGE (%x::'q_40548::type. c) s = COND (s = EMPTY) EMPTY (INSERT c EMPTY)" |
|
| 17644 | 6215 |
by (import hollight IMAGE_CONST) |
6216 |
||
| 19093 | 6217 |
lemma SIMPLE_IMAGE: "ALL (x::'q_40581::type => 'q_40585::type) xa::'q_40581::type => bool. |
| 17644 | 6218 |
GSPEC |
| 19093 | 6219 |
(%u::'q_40585::type. |
6220 |
EX xb::'q_40581::type. SETSPEC u (IN xb xa) (x xb)) = |
|
| 17644 | 6221 |
IMAGE x xa" |
6222 |
by (import hollight SIMPLE_IMAGE) |
|
6223 |
||
| 19093 | 6224 |
lemma EMPTY_GSPEC: "GSPEC (%u::'q_40602::type. Ex (SETSPEC u False)) = EMPTY" |
| 17644 | 6225 |
by (import hollight EMPTY_GSPEC) |
6226 |
||
| 19093 | 6227 |
lemma IN_ELIM_PAIR_THM: "ALL (x::'q_40648::type => 'q_40647::type => bool) (xa::'q_40648::type) |
6228 |
xb::'q_40647::type. |
|
6229 |
IN (xa, xb) |
|
6230 |
(GSPEC |
|
6231 |
(%xa::'q_40648::type * 'q_40647::type. |
|
6232 |
EX (xb::'q_40648::type) y::'q_40647::type. |
|
6233 |
SETSPEC xa (x xb y) (xb, y))) = |
|
6234 |
x xa xb" |
|
6235 |
by (import hollight IN_ELIM_PAIR_THM) |
|
6236 |
||
| 17644 | 6237 |
lemma FINITE_INDUCT_STRONG: "ALL P::('A::type => bool) => bool.
|
6238 |
P EMPTY & |
|
6239 |
(ALL (x::'A::type) s::'A::type => bool. |
|
6240 |
P s & ~ IN x s & FINITE s --> P (INSERT x s)) --> |
|
6241 |
(ALL s::'A::type => bool. FINITE s --> P s)" |
|
6242 |
by (import hollight FINITE_INDUCT_STRONG) |
|
6243 |
||
6244 |
lemma FINITE_SUBSET: "ALL (x::'A::type => bool) t::'A::type => bool. |
|
6245 |
FINITE t & SUBSET x t --> FINITE x" |
|
6246 |
by (import hollight FINITE_SUBSET) |
|
6247 |
||
6248 |
lemma FINITE_UNION_IMP: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
6249 |
FINITE x & FINITE xa --> FINITE (hollight.UNION x xa)" |
|
6250 |
by (import hollight FINITE_UNION_IMP) |
|
6251 |
||
6252 |
lemma FINITE_UNION: "ALL (s::'A::type => bool) t::'A::type => bool. |
|
6253 |
FINITE (hollight.UNION s t) = (FINITE s & FINITE t)" |
|
6254 |
by (import hollight FINITE_UNION) |
|
6255 |
||
6256 |
lemma FINITE_INTER: "ALL (s::'A::type => bool) t::'A::type => bool. |
|
6257 |
FINITE s | FINITE t --> FINITE (hollight.INTER s t)" |
|
6258 |
by (import hollight FINITE_INTER) |
|
6259 |
||
6260 |
lemma FINITE_INSERT: "ALL (s::'A::type => bool) x::'A::type. FINITE (INSERT x s) = FINITE s" |
|
6261 |
by (import hollight FINITE_INSERT) |
|
6262 |
||
6263 |
lemma FINITE_DELETE_IMP: "ALL (s::'A::type => bool) x::'A::type. FINITE s --> FINITE (DELETE s x)" |
|
6264 |
by (import hollight FINITE_DELETE_IMP) |
|
6265 |
||
6266 |
lemma FINITE_DELETE: "ALL (s::'A::type => bool) x::'A::type. FINITE (DELETE s x) = FINITE s" |
|
6267 |
by (import hollight FINITE_DELETE) |
|
6268 |
||
| 19093 | 6269 |
lemma FINITE_UNIONS: "ALL s::('q_40983::type => bool) => bool.
|
| 17644 | 6270 |
FINITE s --> |
| 19093 | 6271 |
FINITE (UNIONS s) = (ALL t::'q_40983::type => bool. IN t s --> FINITE t)" |
| 17644 | 6272 |
by (import hollight FINITE_UNIONS) |
6273 |
||
6274 |
lemma FINITE_IMAGE_EXPAND: "ALL (f::'A::type => 'B::type) s::'A::type => bool. |
|
6275 |
FINITE s --> |
|
6276 |
FINITE |
|
6277 |
(GSPEC |
|
6278 |
(%u::'B::type. |
|
6279 |
EX y::'B::type. SETSPEC u (EX x::'A::type. IN x s & y = f x) y))" |
|
6280 |
by (import hollight FINITE_IMAGE_EXPAND) |
|
6281 |
||
6282 |
lemma FINITE_IMAGE: "ALL (x::'A::type => 'B::type) xa::'A::type => bool. |
|
6283 |
FINITE xa --> FINITE (IMAGE x xa)" |
|
6284 |
by (import hollight FINITE_IMAGE) |
|
6285 |
||
6286 |
lemma FINITE_IMAGE_INJ_GENERAL: "ALL (f::'A::type => 'B::type) (x::'B::type => bool) s::'A::type => bool. |
|
6287 |
(ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y) & |
|
6288 |
FINITE x --> |
|
6289 |
FINITE |
|
6290 |
(GSPEC |
|
6291 |
(%u::'A::type. EX xa::'A::type. SETSPEC u (IN xa s & IN (f xa) x) xa))" |
|
6292 |
by (import hollight FINITE_IMAGE_INJ_GENERAL) |
|
6293 |
||
6294 |
lemma FINITE_IMAGE_INJ: "ALL (f::'A::type => 'B::type) A::'B::type => bool. |
|
6295 |
(ALL (x::'A::type) y::'A::type. f x = f y --> x = y) & FINITE A --> |
|
6296 |
FINITE (GSPEC (%u::'A::type. EX x::'A::type. SETSPEC u (IN (f x) A) x))" |
|
6297 |
by (import hollight FINITE_IMAGE_INJ) |
|
6298 |
||
6299 |
lemma INFINITE_IMAGE_INJ: "ALL f::'A::type => 'B::type. |
|
6300 |
(ALL (x::'A::type) y::'A::type. f x = f y --> x = y) --> |
|
6301 |
(ALL s::'A::type => bool. INFINITE s --> INFINITE (IMAGE f s))" |
|
6302 |
by (import hollight INFINITE_IMAGE_INJ) |
|
6303 |
||
| 19093 | 6304 |
lemma INFINITE_NONEMPTY: "ALL s::'q_41466::type => bool. INFINITE s --> s ~= EMPTY" |
| 17644 | 6305 |
by (import hollight INFINITE_NONEMPTY) |
6306 |
||
6307 |
lemma INFINITE_DIFF_FINITE: "ALL (s::'A::type => bool) t::'A::type => bool. |
|
6308 |
INFINITE s & FINITE t --> INFINITE (DIFF s t)" |
|
6309 |
by (import hollight INFINITE_DIFF_FINITE) |
|
6310 |
||
6311 |
lemma FINITE_SUBSET_IMAGE: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool. |
|
6312 |
(FINITE t & SUBSET t (IMAGE f s)) = |
|
6313 |
(EX x::'A::type => bool. FINITE x & SUBSET x s & t = IMAGE f x)" |
|
6314 |
by (import hollight FINITE_SUBSET_IMAGE) |
|
6315 |
||
6316 |
lemma FINITE_SUBSET_IMAGE_IMP: "ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'B::type => bool. |
|
6317 |
FINITE t & SUBSET t (IMAGE f s) --> |
|
6318 |
(EX s'::'A::type => bool. |
|
6319 |
FINITE s' & SUBSET s' s & SUBSET t (IMAGE f s'))" |
|
6320 |
by (import hollight FINITE_SUBSET_IMAGE_IMP) |
|
6321 |
||
6322 |
lemma FINITE_SUBSETS: "ALL s::'A::type => bool. |
|
6323 |
FINITE s --> |
|
6324 |
FINITE |
|
6325 |
(GSPEC |
|
6326 |
(%u::'A::type => bool. |
|
6327 |
EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))" |
|
6328 |
by (import hollight FINITE_SUBSETS) |
|
6329 |
||
| 19093 | 6330 |
lemma FINITE_DIFF: "ALL (s::'q_41764::type => bool) t::'q_41764::type => bool. |
| 17644 | 6331 |
FINITE s --> FINITE (DIFF s t)" |
6332 |
by (import hollight FINITE_DIFF) |
|
6333 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6334 |
definition FINREC :: "('q_41824 => 'q_41823 => 'q_41823)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6335 |
=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" where
|
| 17644 | 6336 |
"FINREC == |
| 19093 | 6337 |
SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type)
|
6338 |
=> 'q_41823::type |
|
6339 |
=> ('q_41824::type => bool)
|
|
6340 |
=> 'q_41823::type => nat => bool. |
|
6341 |
(ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type) |
|
6342 |
(s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type. |
|
| 17652 | 6343 |
FINREC f b s a 0 = (s = EMPTY & a = b)) & |
| 19093 | 6344 |
(ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat) |
6345 |
(a::'q_41823::type) |
|
6346 |
f::'q_41824::type => 'q_41823::type => 'q_41823::type. |
|
| 17644 | 6347 |
FINREC f b s a (Suc n) = |
| 19093 | 6348 |
(EX (x::'q_41824::type) c::'q_41823::type. |
| 17644 | 6349 |
IN x s & FINREC f b (DELETE s x) c n & a = f x c))" |
6350 |
||
6351 |
lemma DEF_FINREC: "FINREC = |
|
| 19093 | 6352 |
(SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type)
|
6353 |
=> 'q_41823::type |
|
6354 |
=> ('q_41824::type => bool)
|
|
6355 |
=> 'q_41823::type => nat => bool. |
|
6356 |
(ALL (f::'q_41824::type => 'q_41823::type => 'q_41823::type) |
|
6357 |
(s::'q_41824::type => bool) (a::'q_41823::type) b::'q_41823::type. |
|
| 17652 | 6358 |
FINREC f b s a 0 = (s = EMPTY & a = b)) & |
| 19093 | 6359 |
(ALL (b::'q_41823::type) (s::'q_41824::type => bool) (n::nat) |
6360 |
(a::'q_41823::type) |
|
6361 |
f::'q_41824::type => 'q_41823::type => 'q_41823::type. |
|
| 17644 | 6362 |
FINREC f b s a (Suc n) = |
| 19093 | 6363 |
(EX (x::'q_41824::type) c::'q_41823::type. |
| 17644 | 6364 |
IN x s & FINREC f b (DELETE s x) c n & a = f x c)))" |
6365 |
by (import hollight DEF_FINREC) |
|
6366 |
||
| 19093 | 6367 |
lemma FINREC_1_LEMMA: "ALL (x::'q_41869::type => 'q_41868::type => 'q_41868::type) |
6368 |
(xa::'q_41868::type) (xb::'q_41869::type => bool) xc::'q_41868::type. |
|
| 17652 | 6369 |
FINREC x xa xb xc (Suc 0) = |
| 19093 | 6370 |
(EX xd::'q_41869::type. xb = INSERT xd EMPTY & xc = x xd xa)" |
| 17644 | 6371 |
by (import hollight FINREC_1_LEMMA) |
6372 |
||
6373 |
lemma FINREC_SUC_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6374 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6375 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6376 |
(ALL (n::nat) (s::'A::type => bool) z::'B::type. |
|
6377 |
FINREC f b s z (Suc n) --> |
|
6378 |
(ALL x::'A::type. |
|
6379 |
IN x s --> |
|
6380 |
(EX w::'B::type. FINREC f b (DELETE s x) w n & z = f x w)))" |
|
6381 |
by (import hollight FINREC_SUC_LEMMA) |
|
6382 |
||
6383 |
lemma FINREC_UNIQUE_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6384 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6385 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6386 |
(ALL (n1::nat) (n2::nat) (s::'A::type => bool) (a1::'B::type) |
|
6387 |
a2::'B::type. |
|
6388 |
FINREC f b s a1 n1 & FINREC f b s a2 n2 --> a1 = a2 & n1 = n2)" |
|
6389 |
by (import hollight FINREC_UNIQUE_LEMMA) |
|
6390 |
||
6391 |
lemma FINREC_EXISTS_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) (b::'B::type) s::'A::type => bool. |
|
6392 |
FINITE s --> (EX a::'B::type. Ex (FINREC f b s a))" |
|
6393 |
by (import hollight FINREC_EXISTS_LEMMA) |
|
6394 |
||
6395 |
lemma FINREC_FUN_LEMMA: "ALL (P::'A::type => bool) R::'A::type => 'B::type => 'C::type => bool. |
|
6396 |
(ALL s::'A::type. P s --> (EX a::'B::type. Ex (R s a))) & |
|
6397 |
(ALL (n1::'C::type) (n2::'C::type) (s::'A::type) (a1::'B::type) |
|
6398 |
a2::'B::type. R s a1 n1 & R s a2 n2 --> a1 = a2 & n1 = n2) --> |
|
6399 |
(EX x::'A::type => 'B::type. |
|
6400 |
ALL (s::'A::type) a::'B::type. P s --> Ex (R s a) = (x s = a))" |
|
6401 |
by (import hollight FINREC_FUN_LEMMA) |
|
6402 |
||
6403 |
lemma FINREC_FUN: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6404 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6405 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6406 |
(EX g::('A::type => bool) => 'B::type.
|
|
6407 |
g EMPTY = b & |
|
6408 |
(ALL (s::'A::type => bool) x::'A::type. |
|
6409 |
FINITE s & IN x s --> g s = f x (g (DELETE s x))))" |
|
6410 |
by (import hollight FINREC_FUN) |
|
6411 |
||
6412 |
lemma SET_RECURSION_LEMMA: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6413 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6414 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6415 |
(EX g::('A::type => bool) => 'B::type.
|
|
6416 |
g EMPTY = b & |
|
6417 |
(ALL (x::'A::type) s::'A::type => bool. |
|
6418 |
FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (f x (g s))))" |
|
6419 |
by (import hollight SET_RECURSION_LEMMA) |
|
6420 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6421 |
definition ITSET :: "('q_42525 => 'q_42524 => 'q_42524)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6422 |
=> ('q_42525 => bool) => 'q_42524 => 'q_42524" where
|
| 17644 | 6423 |
"ITSET == |
| 19093 | 6424 |
%(u::'q_42525::type => 'q_42524::type => 'q_42524::type) |
6425 |
(ua::'q_42525::type => bool) ub::'q_42524::type. |
|
6426 |
(SOME g::('q_42525::type => bool) => 'q_42524::type.
|
|
| 17644 | 6427 |
g EMPTY = ub & |
| 19093 | 6428 |
(ALL (x::'q_42525::type) s::'q_42525::type => bool. |
| 17644 | 6429 |
FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s)))) |
6430 |
ua" |
|
6431 |
||
6432 |
lemma DEF_ITSET: "ITSET = |
|
| 19093 | 6433 |
(%(u::'q_42525::type => 'q_42524::type => 'q_42524::type) |
6434 |
(ua::'q_42525::type => bool) ub::'q_42524::type. |
|
6435 |
(SOME g::('q_42525::type => bool) => 'q_42524::type.
|
|
| 17644 | 6436 |
g EMPTY = ub & |
| 19093 | 6437 |
(ALL (x::'q_42525::type) s::'q_42525::type => bool. |
| 17644 | 6438 |
FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (u x (g s)))) |
6439 |
ua)" |
|
6440 |
by (import hollight DEF_ITSET) |
|
6441 |
||
6442 |
lemma FINITE_RECURSION: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6443 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6444 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6445 |
ITSET f EMPTY b = b & |
|
6446 |
(ALL (x::'A::type) xa::'A::type => bool. |
|
6447 |
FINITE xa --> |
|
6448 |
ITSET f (INSERT x xa) b = |
|
6449 |
COND (IN x xa) (ITSET f xa b) (f x (ITSET f xa b)))" |
|
6450 |
by (import hollight FINITE_RECURSION) |
|
6451 |
||
6452 |
lemma FINITE_RECURSION_DELETE: "ALL (f::'A::type => 'B::type => 'B::type) b::'B::type. |
|
6453 |
(ALL (x::'A::type) (y::'A::type) s::'B::type. |
|
6454 |
x ~= y --> f x (f y s) = f y (f x s)) --> |
|
6455 |
ITSET f EMPTY b = b & |
|
6456 |
(ALL (x::'A::type) s::'A::type => bool. |
|
6457 |
FINITE s --> |
|
6458 |
ITSET f s b = |
|
6459 |
COND (IN x s) (f x (ITSET f (DELETE s x) b)) |
|
6460 |
(ITSET f (DELETE s x) b))" |
|
6461 |
by (import hollight FINITE_RECURSION_DELETE) |
|
6462 |
||
| 19093 | 6463 |
lemma ITSET_EQ: "ALL (x::'q_42830::type => bool) |
6464 |
(xa::'q_42830::type => 'q_42831::type => 'q_42831::type) |
|
6465 |
(xb::'q_42830::type => 'q_42831::type => 'q_42831::type) |
|
6466 |
xc::'q_42831::type. |
|
| 17644 | 6467 |
FINITE x & |
| 19093 | 6468 |
(ALL xc::'q_42830::type. IN xc x --> xa xc = xb xc) & |
6469 |
(ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type. |
|
| 17644 | 6470 |
x ~= y --> xa x (xa y s) = xa y (xa x s)) & |
| 19093 | 6471 |
(ALL (x::'q_42830::type) (y::'q_42830::type) s::'q_42831::type. |
| 17644 | 6472 |
x ~= y --> xb x (xb y s) = xb y (xb x s)) --> |
6473 |
ITSET xa x xc = ITSET xb x xc" |
|
6474 |
by (import hollight ITSET_EQ) |
|
6475 |
||
| 19093 | 6476 |
lemma SUBSET_RESTRICT: "ALL (x::'q_42864::type => bool) xa::'q_42864::type => bool. |
| 17644 | 6477 |
SUBSET |
6478 |
(GSPEC |
|
| 19093 | 6479 |
(%u::'q_42864::type. |
6480 |
EX xb::'q_42864::type. SETSPEC u (IN xb x & xa xb) xb)) |
|
| 17644 | 6481 |
x" |
6482 |
by (import hollight SUBSET_RESTRICT) |
|
6483 |
||
| 19093 | 6484 |
lemma FINITE_RESTRICT: "ALL (s::'A::type => bool) p::'q_42882::type. |
| 17644 | 6485 |
FINITE s --> |
6486 |
FINITE |
|
6487 |
(GSPEC |
|
6488 |
(%u::'A::type. |
|
6489 |
EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))" |
|
6490 |
by (import hollight FINITE_RESTRICT) |
|
6491 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6492 |
definition CARD :: "('q_42918 => bool) => nat" where
|
| 19093 | 6493 |
"CARD == %u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0" |
6494 |
||
6495 |
lemma DEF_CARD: "CARD = (%u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0)" |
|
| 17644 | 6496 |
by (import hollight DEF_CARD) |
6497 |
||
6498 |
lemma CARD_CLAUSES: "(op &::bool => bool => bool) |
|
6499 |
((op =::nat => nat => bool) |
|
6500 |
((CARD::('A::type => bool) => nat) (EMPTY::'A::type => bool)) (0::nat))
|
|
6501 |
((All::('A::type => bool) => bool)
|
|
6502 |
(%x::'A::type. |
|
6503 |
(All::(('A::type => bool) => bool) => bool)
|
|
6504 |
(%s::'A::type => bool. |
|
6505 |
(op -->::bool => bool => bool) |
|
6506 |
((FINITE::('A::type => bool) => bool) s)
|
|
6507 |
((op =::nat => nat => bool) |
|
6508 |
((CARD::('A::type => bool) => nat)
|
|
6509 |
((INSERT::'A::type |
|
6510 |
=> ('A::type => bool) => 'A::type => bool)
|
|
6511 |
x s)) |
|
6512 |
((COND::bool => nat => nat => nat) |
|
6513 |
((IN::'A::type => ('A::type => bool) => bool) x s)
|
|
6514 |
((CARD::('A::type => bool) => nat) s)
|
|
6515 |
((Suc::nat => nat) |
|
6516 |
((CARD::('A::type => bool) => nat) s)))))))"
|
|
6517 |
by (import hollight CARD_CLAUSES) |
|
6518 |
||
6519 |
lemma CARD_UNION: "ALL (x::'A::type => bool) xa::'A::type => bool. |
|
6520 |
FINITE x & FINITE xa & hollight.INTER x xa = EMPTY --> |
|
6521 |
CARD (hollight.UNION x xa) = CARD x + CARD xa" |
|
6522 |
by (import hollight CARD_UNION) |
|
6523 |
||
6524 |
lemma CARD_DELETE: "ALL (x::'A::type) s::'A::type => bool. |
|
6525 |
FINITE s --> |
|
| 17652 | 6526 |
CARD (DELETE s x) = COND (IN x s) (CARD s - NUMERAL_BIT1 0) (CARD s)" |
| 17644 | 6527 |
by (import hollight CARD_DELETE) |
6528 |
||
| 19093 | 6529 |
lemma CARD_UNION_EQ: "ALL (s::'q_43163::type => bool) (t::'q_43163::type => bool) |
6530 |
u::'q_43163::type => bool. |
|
| 17644 | 6531 |
FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> |
6532 |
CARD s + CARD t = CARD u" |
|
6533 |
by (import hollight CARD_UNION_EQ) |
|
6534 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6535 |
definition HAS_SIZE :: "('q_43199 => bool) => nat => bool" where
|
| 19093 | 6536 |
"HAS_SIZE == %(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua" |
6537 |
||
6538 |
lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua)" |
|
| 17644 | 6539 |
by (import hollight DEF_HAS_SIZE) |
6540 |
||
| 19093 | 6541 |
lemma HAS_SIZE_CARD: "ALL (x::'q_43218::type => bool) xa::nat. HAS_SIZE x xa --> CARD x = xa" |
| 17644 | 6542 |
by (import hollight HAS_SIZE_CARD) |
6543 |
||
| 19093 | 6544 |
lemma HAS_SIZE_0: "ALL (s::'A::type => bool) n::'q_43234::type. HAS_SIZE s 0 = (s = EMPTY)" |
| 17644 | 6545 |
by (import hollight HAS_SIZE_0) |
6546 |
||
6547 |
lemma HAS_SIZE_SUC: "ALL (s::'A::type => bool) n::nat. |
|
6548 |
HAS_SIZE s (Suc n) = |
|
6549 |
(s ~= EMPTY & (ALL x::'A::type. IN x s --> HAS_SIZE (DELETE s x) n))" |
|
6550 |
by (import hollight HAS_SIZE_SUC) |
|
6551 |
||
| 19093 | 6552 |
lemma HAS_SIZE_UNION: "ALL (x::'q_43356::type => bool) (xa::'q_43356::type => bool) (xb::nat) |
| 17644 | 6553 |
xc::nat. |
6554 |
HAS_SIZE x xb & HAS_SIZE xa xc & DISJOINT x xa --> |
|
6555 |
HAS_SIZE (hollight.UNION x xa) (xb + xc)" |
|
6556 |
by (import hollight HAS_SIZE_UNION) |
|
6557 |
||
6558 |
lemma HAS_SIZE_UNIONS: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat) |
|
6559 |
xc::nat. |
|
6560 |
HAS_SIZE x xb & |
|
6561 |
(ALL xb::'A::type. IN xb x --> HAS_SIZE (xa xb) xc) & |
|
6562 |
(ALL (xb::'A::type) y::'A::type. |
|
6563 |
IN xb x & IN y x & xb ~= y --> DISJOINT (xa xb) (xa y)) --> |
|
6564 |
HAS_SIZE |
|
6565 |
(UNIONS |
|
6566 |
(GSPEC |
|
6567 |
(%u::'B::type => bool. |
|
6568 |
EX xb::'A::type. SETSPEC u (IN xb x) (xa xb)))) |
|
6569 |
(xb * xc)" |
|
6570 |
by (import hollight HAS_SIZE_UNIONS) |
|
6571 |
||
| 19093 | 6572 |
lemma HAS_SIZE_CLAUSES: "HAS_SIZE (s::'q_43604::type => bool) 0 = (s = EMPTY) & |
| 17644 | 6573 |
HAS_SIZE s (Suc (n::nat)) = |
| 19093 | 6574 |
(EX (a::'q_43604::type) t::'q_43604::type => bool. |
| 17644 | 6575 |
HAS_SIZE t n & ~ IN a t & s = INSERT a t)" |
6576 |
by (import hollight HAS_SIZE_CLAUSES) |
|
6577 |
||
6578 |
lemma CARD_SUBSET_EQ: "ALL (a::'A::type => bool) b::'A::type => bool. |
|
6579 |
FINITE b & SUBSET a b & CARD a = CARD b --> a = b" |
|
6580 |
by (import hollight CARD_SUBSET_EQ) |
|
6581 |
||
6582 |
lemma CARD_SUBSET: "ALL (a::'A::type => bool) b::'A::type => bool. |
|
6583 |
SUBSET a b & FINITE b --> <= (CARD a) (CARD b)" |
|
6584 |
by (import hollight CARD_SUBSET) |
|
6585 |
||
6586 |
lemma CARD_SUBSET_LE: "ALL (a::'A::type => bool) b::'A::type => bool. |
|
6587 |
FINITE b & SUBSET a b & <= (CARD b) (CARD a) --> a = b" |
|
6588 |
by (import hollight CARD_SUBSET_LE) |
|
6589 |
||
| 19093 | 6590 |
lemma CARD_EQ_0: "ALL s::'q_43920::type => bool. FINITE s --> (CARD s = 0) = (s = EMPTY)" |
| 17644 | 6591 |
by (import hollight CARD_EQ_0) |
6592 |
||
6593 |
lemma CARD_PSUBSET: "ALL (a::'A::type => bool) b::'A::type => bool. |
|
6594 |
PSUBSET a b & FINITE b --> < (CARD a) (CARD b)" |
|
6595 |
by (import hollight CARD_PSUBSET) |
|
6596 |
||
6597 |
lemma CARD_UNION_LE: "ALL (s::'A::type => bool) t::'A::type => bool. |
|
6598 |
FINITE s & FINITE t --> <= (CARD (hollight.UNION s t)) (CARD s + CARD t)" |
|
6599 |
by (import hollight CARD_UNION_LE) |
|
6600 |
||
6601 |
lemma CARD_UNIONS_LE: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) (xb::nat) |
|
6602 |
xc::nat. |
|
6603 |
HAS_SIZE x xb & |
|
6604 |
(ALL xb::'A::type. IN xb x --> FINITE (xa xb) & <= (CARD (xa xb)) xc) --> |
|
6605 |
<= (CARD |
|
6606 |
(UNIONS |
|
6607 |
(GSPEC |
|
6608 |
(%u::'B::type => bool. |
|
6609 |
EX xb::'A::type. SETSPEC u (IN xb x) (xa xb))))) |
|
6610 |
(xb * xc)" |
|
6611 |
by (import hollight CARD_UNIONS_LE) |
|
6612 |
||
6613 |
lemma CARD_IMAGE_INJ: "ALL (f::'A::type => 'B::type) x::'A::type => bool. |
|
6614 |
(ALL (xa::'A::type) y::'A::type. |
|
6615 |
IN xa x & IN y x & f xa = f y --> xa = y) & |
|
6616 |
FINITE x --> |
|
6617 |
CARD (IMAGE f x) = CARD x" |
|
6618 |
by (import hollight CARD_IMAGE_INJ) |
|
6619 |
||
6620 |
lemma HAS_SIZE_IMAGE_INJ: "ALL (x::'A::type => 'B::type) (xa::'A::type => bool) xb::nat. |
|
6621 |
(ALL (xb::'A::type) y::'A::type. |
|
6622 |
IN xb xa & IN y xa & x xb = x y --> xb = y) & |
|
6623 |
HAS_SIZE xa xb --> |
|
6624 |
HAS_SIZE (IMAGE x xa) xb" |
|
6625 |
by (import hollight HAS_SIZE_IMAGE_INJ) |
|
6626 |
||
6627 |
lemma CARD_IMAGE_LE: "ALL (f::'A::type => 'B::type) s::'A::type => bool. |
|
6628 |
FINITE s --> <= (CARD (IMAGE f s)) (CARD s)" |
|
6629 |
by (import hollight CARD_IMAGE_LE) |
|
6630 |
||
6631 |
lemma CHOOSE_SUBSET: "ALL s::'A::type => bool. |
|
6632 |
FINITE s --> |
|
6633 |
(ALL n::nat. |
|
6634 |
<= n (CARD s) --> |
|
6635 |
(EX t::'A::type => bool. SUBSET t s & HAS_SIZE t n))" |
|
6636 |
by (import hollight CHOOSE_SUBSET) |
|
6637 |
||
6638 |
lemma HAS_SIZE_PRODUCT_DEPENDENT: "ALL (x::'A::type => bool) (xa::nat) (xb::'A::type => 'B::type => bool) |
|
6639 |
xc::nat. |
|
6640 |
HAS_SIZE x xa & (ALL xa::'A::type. IN xa x --> HAS_SIZE (xb xa) xc) --> |
|
6641 |
HAS_SIZE |
|
6642 |
(GSPEC |
|
6643 |
(%u::'A::type * 'B::type. |
|
6644 |
EX (xa::'A::type) y::'B::type. |
|
6645 |
SETSPEC u (IN xa x & IN y (xb xa)) (xa, y))) |
|
6646 |
(xa * xc)" |
|
6647 |
by (import hollight HAS_SIZE_PRODUCT_DEPENDENT) |
|
6648 |
||
6649 |
lemma FINITE_PRODUCT_DEPENDENT: "ALL (x::'A::type => bool) xa::'A::type => 'B::type => bool. |
|
6650 |
FINITE x & (ALL xb::'A::type. IN xb x --> FINITE (xa xb)) --> |
|
6651 |
FINITE |
|
6652 |
(GSPEC |
|
6653 |
(%u::'A::type * 'B::type. |
|
6654 |
EX (xb::'A::type) y::'B::type. |
|
6655 |
SETSPEC u (IN xb x & IN y (xa xb)) (xb, y)))" |
|
6656 |
by (import hollight FINITE_PRODUCT_DEPENDENT) |
|
6657 |
||
6658 |
lemma FINITE_PRODUCT: "ALL (x::'A::type => bool) xa::'B::type => bool. |
|
6659 |
FINITE x & FINITE xa --> |
|
6660 |
FINITE |
|
6661 |
(GSPEC |
|
6662 |
(%u::'A::type * 'B::type. |
|
6663 |
EX (xb::'A::type) y::'B::type. |
|
6664 |
SETSPEC u (IN xb x & IN y xa) (xb, y)))" |
|
6665 |
by (import hollight FINITE_PRODUCT) |
|
6666 |
||
6667 |
lemma CARD_PRODUCT: "ALL (s::'A::type => bool) t::'B::type => bool. |
|
6668 |
FINITE s & FINITE t --> |
|
6669 |
CARD |
|
6670 |
(GSPEC |
|
6671 |
(%u::'A::type * 'B::type. |
|
6672 |
EX (x::'A::type) y::'B::type. |
|
6673 |
SETSPEC u (IN x s & IN y t) (x, y))) = |
|
6674 |
CARD s * CARD t" |
|
6675 |
by (import hollight CARD_PRODUCT) |
|
6676 |
||
6677 |
lemma HAS_SIZE_PRODUCT: "ALL (x::'A::type => bool) (xa::nat) (xb::'B::type => bool) xc::nat. |
|
6678 |
HAS_SIZE x xa & HAS_SIZE xb xc --> |
|
6679 |
HAS_SIZE |
|
6680 |
(GSPEC |
|
6681 |
(%u::'A::type * 'B::type. |
|
6682 |
EX (xa::'A::type) y::'B::type. |
|
6683 |
SETSPEC u (IN xa x & IN y xb) (xa, y))) |
|
6684 |
(xa * xc)" |
|
6685 |
by (import hollight HAS_SIZE_PRODUCT) |
|
6686 |
||
6687 |
lemma HAS_SIZE_FUNSPACE: "ALL (d::'B::type) (n::nat) (t::'B::type => bool) (m::nat) |
|
6688 |
s::'A::type => bool. |
|
6689 |
HAS_SIZE s m & HAS_SIZE t n --> |
|
6690 |
HAS_SIZE |
|
6691 |
(GSPEC |
|
6692 |
(%u::'A::type => 'B::type. |
|
6693 |
EX f::'A::type => 'B::type. |
|
6694 |
SETSPEC u |
|
6695 |
((ALL x::'A::type. IN x s --> IN (f x) t) & |
|
6696 |
(ALL x::'A::type. ~ IN x s --> f x = d)) |
|
6697 |
f)) |
|
6698 |
(EXP n m)" |
|
6699 |
by (import hollight HAS_SIZE_FUNSPACE) |
|
6700 |
||
| 19093 | 6701 |
lemma CARD_FUNSPACE: "ALL (s::'q_45275::type => bool) t::'q_45272::type => bool. |
| 17644 | 6702 |
FINITE s & FINITE t --> |
6703 |
CARD |
|
6704 |
(GSPEC |
|
| 19093 | 6705 |
(%u::'q_45275::type => 'q_45272::type. |
6706 |
EX f::'q_45275::type => 'q_45272::type. |
|
| 17644 | 6707 |
SETSPEC u |
| 19093 | 6708 |
((ALL x::'q_45275::type. IN x s --> IN (f x) t) & |
6709 |
(ALL x::'q_45275::type. |
|
6710 |
~ IN x s --> f x = (d::'q_45272::type))) |
|
| 17644 | 6711 |
f)) = |
6712 |
EXP (CARD t) (CARD s)" |
|
6713 |
by (import hollight CARD_FUNSPACE) |
|
6714 |
||
| 19093 | 6715 |
lemma FINITE_FUNSPACE: "ALL (s::'q_45341::type => bool) t::'q_45338::type => bool. |
| 17644 | 6716 |
FINITE s & FINITE t --> |
6717 |
FINITE |
|
6718 |
(GSPEC |
|
| 19093 | 6719 |
(%u::'q_45341::type => 'q_45338::type. |
6720 |
EX f::'q_45341::type => 'q_45338::type. |
|
| 17644 | 6721 |
SETSPEC u |
| 19093 | 6722 |
((ALL x::'q_45341::type. IN x s --> IN (f x) t) & |
6723 |
(ALL x::'q_45341::type. |
|
6724 |
~ IN x s --> f x = (d::'q_45338::type))) |
|
| 17644 | 6725 |
f))" |
6726 |
by (import hollight FINITE_FUNSPACE) |
|
6727 |
||
6728 |
lemma HAS_SIZE_POWERSET: "ALL (s::'A::type => bool) n::nat. |
|
6729 |
HAS_SIZE s n --> |
|
6730 |
HAS_SIZE |
|
6731 |
(GSPEC |
|
6732 |
(%u::'A::type => bool. |
|
6733 |
EX t::'A::type => bool. SETSPEC u (SUBSET t s) t)) |
|
| 17652 | 6734 |
(EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) n)" |
| 17644 | 6735 |
by (import hollight HAS_SIZE_POWERSET) |
6736 |
||
6737 |
lemma CARD_POWERSET: "ALL s::'A::type => bool. |
|
6738 |
FINITE s --> |
|
6739 |
CARD |
|
6740 |
(GSPEC |
|
6741 |
(%u::'A::type => bool. |
|
6742 |
EX t::'A::type => bool. SETSPEC u (SUBSET t s) t)) = |
|
| 17652 | 6743 |
EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) (CARD s)" |
| 17644 | 6744 |
by (import hollight CARD_POWERSET) |
6745 |
||
6746 |
lemma FINITE_POWERSET: "ALL s::'A::type => bool. |
|
6747 |
FINITE s --> |
|
6748 |
FINITE |
|
6749 |
(GSPEC |
|
6750 |
(%u::'A::type => bool. |
|
6751 |
EX t::'A::type => bool. SETSPEC u (SUBSET t s) t))" |
|
6752 |
by (import hollight FINITE_POWERSET) |
|
6753 |
||
6754 |
lemma CARD_GE_REFL: "ALL s::'A::type => bool. CARD_GE s s" |
|
6755 |
by (import hollight CARD_GE_REFL) |
|
6756 |
||
6757 |
lemma CARD_GE_TRANS: "ALL (s::'A::type => bool) (t::'B::type => bool) u::'C::type => bool. |
|
6758 |
CARD_GE s t & CARD_GE t u --> CARD_GE s u" |
|
6759 |
by (import hollight CARD_GE_TRANS) |
|
6760 |
||
6761 |
lemma FINITE_HAS_SIZE_LEMMA: "ALL s::'A::type => bool. |
|
6762 |
FINITE s --> |
|
6763 |
(EX n::nat. CARD_GE (GSPEC (%u::nat. EX x::nat. SETSPEC u (< x n) x)) s)" |
|
6764 |
by (import hollight FINITE_HAS_SIZE_LEMMA) |
|
6765 |
||
6766 |
lemma HAS_SIZE_NUMSEG_LT: "ALL n::nat. HAS_SIZE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m n) m)) n" |
|
6767 |
by (import hollight HAS_SIZE_NUMSEG_LT) |
|
6768 |
||
6769 |
lemma CARD_NUMSEG_LT: "ALL x::nat. CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m)) = x" |
|
6770 |
by (import hollight CARD_NUMSEG_LT) |
|
6771 |
||
6772 |
lemma FINITE_NUMSEG_LT: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (< m x) m))" |
|
6773 |
by (import hollight FINITE_NUMSEG_LT) |
|
6774 |
||
6775 |
lemma HAS_SIZE_NUMSEG_LE: "ALL x::nat. |
|
6776 |
HAS_SIZE (GSPEC (%xa::nat. EX xb::nat. SETSPEC xa (<= xb x) xb)) |
|
| 17652 | 6777 |
(x + NUMERAL_BIT1 0)" |
| 17644 | 6778 |
by (import hollight HAS_SIZE_NUMSEG_LE) |
6779 |
||
6780 |
lemma FINITE_NUMSEG_LE: "ALL x::nat. FINITE (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m))" |
|
6781 |
by (import hollight FINITE_NUMSEG_LE) |
|
6782 |
||
6783 |
lemma CARD_NUMSEG_LE: "ALL x::nat. |
|
6784 |
CARD (GSPEC (%u::nat. EX m::nat. SETSPEC u (<= m x) m)) = |
|
| 17652 | 6785 |
x + NUMERAL_BIT1 0" |
| 17644 | 6786 |
by (import hollight CARD_NUMSEG_LE) |
6787 |
||
6788 |
lemma num_FINITE: "ALL s::nat => bool. FINITE s = (EX a::nat. ALL x::nat. IN x s --> <= x a)" |
|
6789 |
by (import hollight num_FINITE) |
|
6790 |
||
6791 |
lemma num_FINITE_AVOID: "ALL s::nat => bool. FINITE s --> (EX a::nat. ~ IN a s)" |
|
6792 |
by (import hollight num_FINITE_AVOID) |
|
6793 |
||
6794 |
lemma num_INFINITE: "(INFINITE::(nat => bool) => bool) (hollight.UNIV::nat => bool)" |
|
6795 |
by (import hollight num_INFINITE) |
|
6796 |
||
6797 |
lemma HAS_SIZE_INDEX: "ALL (x::'A::type => bool) n::nat. |
|
6798 |
HAS_SIZE x n --> |
|
6799 |
(EX f::nat => 'A::type. |
|
6800 |
(ALL m::nat. < m n --> IN (f m) x) & |
|
6801 |
(ALL xa::'A::type. IN xa x --> (EX! m::nat. < m n & f m = xa)))" |
|
6802 |
by (import hollight HAS_SIZE_INDEX) |
|
6803 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6804 |
definition set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" where |
| 17644 | 6805 |
"set_of_list == |
| 19093 | 6806 |
SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool. |
| 17644 | 6807 |
set_of_list NIL = EMPTY & |
| 19093 | 6808 |
(ALL (h::'q_45968::type) t::'q_45968::type hollight.list. |
| 17644 | 6809 |
set_of_list (CONS h t) = INSERT h (set_of_list t))" |
6810 |
||
6811 |
lemma DEF_set_of_list: "set_of_list = |
|
| 19093 | 6812 |
(SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool. |
| 17644 | 6813 |
set_of_list NIL = EMPTY & |
| 19093 | 6814 |
(ALL (h::'q_45968::type) t::'q_45968::type hollight.list. |
| 17644 | 6815 |
set_of_list (CONS h t) = INSERT h (set_of_list t)))" |
6816 |
by (import hollight DEF_set_of_list) |
|
6817 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6818 |
definition list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" where
|
| 17644 | 6819 |
"list_of_set == |
| 19093 | 6820 |
%u::'q_45986::type => bool. |
6821 |
SOME l::'q_45986::type hollight.list. |
|
| 17644 | 6822 |
set_of_list l = u & LENGTH l = CARD u" |
6823 |
||
6824 |
lemma DEF_list_of_set: "list_of_set = |
|
| 19093 | 6825 |
(%u::'q_45986::type => bool. |
6826 |
SOME l::'q_45986::type hollight.list. |
|
| 17644 | 6827 |
set_of_list l = u & LENGTH l = CARD u)" |
6828 |
by (import hollight DEF_list_of_set) |
|
6829 |
||
6830 |
lemma LIST_OF_SET_PROPERTIES: "ALL x::'A::type => bool. |
|
6831 |
FINITE x --> |
|
6832 |
set_of_list (list_of_set x) = x & LENGTH (list_of_set x) = CARD x" |
|
6833 |
by (import hollight LIST_OF_SET_PROPERTIES) |
|
6834 |
||
| 19093 | 6835 |
lemma SET_OF_LIST_OF_SET: "ALL s::'q_46035::type => bool. FINITE s --> set_of_list (list_of_set s) = s" |
| 17644 | 6836 |
by (import hollight SET_OF_LIST_OF_SET) |
6837 |
||
| 19093 | 6838 |
lemma LENGTH_LIST_OF_SET: "ALL s::'q_46051::type => bool. FINITE s --> LENGTH (list_of_set s) = CARD s" |
| 17644 | 6839 |
by (import hollight LENGTH_LIST_OF_SET) |
6840 |
||
6841 |
lemma MEM_LIST_OF_SET: "ALL s::'A::type => bool. |
|
6842 |
FINITE s --> (ALL x::'A::type. MEM x (list_of_set s) = IN x s)" |
|
6843 |
by (import hollight MEM_LIST_OF_SET) |
|
6844 |
||
| 19093 | 6845 |
lemma FINITE_SET_OF_LIST: "ALL l::'q_46096::type hollight.list. FINITE (set_of_list l)" |
| 17644 | 6846 |
by (import hollight FINITE_SET_OF_LIST) |
6847 |
||
| 19093 | 6848 |
lemma IN_SET_OF_LIST: "ALL (x::'q_46114::type) l::'q_46114::type hollight.list. |
| 17644 | 6849 |
IN x (set_of_list l) = MEM x l" |
6850 |
by (import hollight IN_SET_OF_LIST) |
|
6851 |
||
| 19093 | 6852 |
lemma SET_OF_LIST_APPEND: "ALL (x::'q_46139::type hollight.list) xa::'q_46139::type hollight.list. |
| 17644 | 6853 |
set_of_list (APPEND x xa) = |
6854 |
hollight.UNION (set_of_list x) (set_of_list xa)" |
|
6855 |
by (import hollight SET_OF_LIST_APPEND) |
|
6856 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6857 |
definition pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" where
|
| 17644 | 6858 |
"pairwise == |
| 19093 | 6859 |
%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool. |
6860 |
ALL (x::'q_46198::type) y::'q_46198::type. |
|
| 17644 | 6861 |
IN x ua & IN y ua & x ~= y --> u x y" |
6862 |
||
6863 |
lemma DEF_pairwise: "pairwise = |
|
| 19093 | 6864 |
(%(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool. |
6865 |
ALL (x::'q_46198::type) y::'q_46198::type. |
|
| 17644 | 6866 |
IN x ua & IN y ua & x ~= y --> u x y)" |
6867 |
by (import hollight DEF_pairwise) |
|
6868 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6869 |
definition PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" where
|
| 17644 | 6870 |
"PAIRWISE == |
| 19093 | 6871 |
SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool)
|
6872 |
=> 'q_46220::type hollight.list => bool. |
|
6873 |
(ALL r::'q_46220::type => 'q_46220::type => bool. |
|
| 17644 | 6874 |
PAIRWISE r NIL = True) & |
| 19093 | 6875 |
(ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool) |
6876 |
t::'q_46220::type hollight.list. |
|
| 17644 | 6877 |
PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t))" |
6878 |
||
6879 |
lemma DEF_PAIRWISE: "PAIRWISE = |
|
| 19093 | 6880 |
(SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool)
|
6881 |
=> 'q_46220::type hollight.list => bool. |
|
6882 |
(ALL r::'q_46220::type => 'q_46220::type => bool. |
|
| 17644 | 6883 |
PAIRWISE r NIL = True) & |
| 19093 | 6884 |
(ALL (h::'q_46220::type) (r::'q_46220::type => 'q_46220::type => bool) |
6885 |
t::'q_46220::type hollight.list. |
|
| 17644 | 6886 |
PAIRWISE r (CONS h t) = (ALL_list (r h) t & PAIRWISE r t)))" |
6887 |
by (import hollight DEF_PAIRWISE) |
|
6888 |
||
6889 |
typedef (open) ('A) finite_image = "(Collect::('A::type => bool) => 'A::type set)
|
|
6890 |
(%x::'A::type. |
|
6891 |
(op |::bool => bool => bool) |
|
6892 |
((op =::'A::type => 'A::type => bool) x |
|
6893 |
((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool)))
|
|
6894 |
((FINITE::('A::type => bool) => bool)
|
|
6895 |
(hollight.UNIV::'A::type => bool)))" morphisms "dest_finite_image" "mk_finite_image" |
|
6896 |
apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type => bool) => 'A::type)
|
|
6897 |
(%x::'A::type. |
|
6898 |
(op |::bool => bool => bool) |
|
6899 |
((op =::'A::type => 'A::type => bool) x |
|
6900 |
((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool)))
|
|
6901 |
((FINITE::('A::type => bool) => bool)
|
|
6902 |
(hollight.UNIV::'A::type => bool)))"]) |
|
6903 |
by (import hollight TYDEF_finite_image) |
|
6904 |
||
6905 |
syntax |
|
6906 |
dest_finite_image :: _ |
|
6907 |
||
6908 |
syntax |
|
6909 |
mk_finite_image :: _ |
|
6910 |
||
6911 |
lemmas "TYDEF_finite_image_@intern" = typedef_hol2hollight |
|
| 17652 | 6912 |
[where a="a :: 'A finite_image" and r=r , |
| 17644 | 6913 |
OF type_definition_finite_image] |
6914 |
||
6915 |
lemma FINITE_IMAGE_IMAGE: "(op =::('A::type finite_image => bool)
|
|
6916 |
=> ('A::type finite_image => bool) => bool)
|
|
6917 |
(hollight.UNIV::'A::type finite_image => bool) |
|
6918 |
((IMAGE::('A::type => 'A::type finite_image)
|
|
6919 |
=> ('A::type => bool) => 'A::type finite_image => bool)
|
|
6920 |
(mk_finite_image::'A::type => 'A::type finite_image) |
|
6921 |
((COND::bool |
|
6922 |
=> ('A::type => bool) => ('A::type => bool) => 'A::type => bool)
|
|
6923 |
((FINITE::('A::type => bool) => bool)
|
|
6924 |
(hollight.UNIV::'A::type => bool)) |
|
6925 |
(hollight.UNIV::'A::type => bool) |
|
6926 |
((INSERT::'A::type => ('A::type => bool) => 'A::type => bool)
|
|
6927 |
((Eps::('A::type => bool) => 'A::type) (%z::'A::type. True::bool))
|
|
6928 |
(EMPTY::'A::type => bool))))" |
|
6929 |
by (import hollight FINITE_IMAGE_IMAGE) |
|
6930 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6931 |
definition dimindex :: "('A => bool) => nat" where
|
| 17644 | 6932 |
"(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop)
|
6933 |
(dimindex::('A::type => bool) => nat)
|
|
6934 |
(%u::'A::type => bool. |
|
6935 |
(COND::bool => nat => nat => nat) |
|
6936 |
((FINITE::('A::type => bool) => bool)
|
|
6937 |
(hollight.UNIV::'A::type => bool)) |
|
6938 |
((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool))
|
|
6939 |
((NUMERAL_BIT1::nat => nat) (0::nat)))" |
|
6940 |
||
6941 |
lemma DEF_dimindex: "(op =::(('A::type => bool) => nat) => (('A::type => bool) => nat) => bool)
|
|
6942 |
(dimindex::('A::type => bool) => nat)
|
|
6943 |
(%u::'A::type => bool. |
|
6944 |
(COND::bool => nat => nat => nat) |
|
6945 |
((FINITE::('A::type => bool) => bool)
|
|
6946 |
(hollight.UNIV::'A::type => bool)) |
|
6947 |
((CARD::('A::type => bool) => nat) (hollight.UNIV::'A::type => bool))
|
|
6948 |
((NUMERAL_BIT1::nat => nat) (0::nat)))" |
|
6949 |
by (import hollight DEF_dimindex) |
|
6950 |
||
6951 |
lemma HAS_SIZE_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool)
|
|
6952 |
(%s::'A::type => bool. |
|
6953 |
(HAS_SIZE::('A::type finite_image => bool) => nat => bool)
|
|
6954 |
(hollight.UNIV::'A::type finite_image => bool) |
|
6955 |
((dimindex::('A::type => bool) => nat) s))"
|
|
6956 |
by (import hollight HAS_SIZE_FINITE_IMAGE) |
|
6957 |
||
6958 |
lemma CARD_FINITE_IMAGE: "(All::(('A::type => bool) => bool) => bool)
|
|
6959 |
(%s::'A::type => bool. |
|
6960 |
(op =::nat => nat => bool) |
|
6961 |
((CARD::('A::type finite_image => bool) => nat)
|
|
6962 |
(hollight.UNIV::'A::type finite_image => bool)) |
|
6963 |
((dimindex::('A::type => bool) => nat) s))"
|
|
6964 |
by (import hollight CARD_FINITE_IMAGE) |
|
6965 |
||
6966 |
lemma FINITE_FINITE_IMAGE: "(FINITE::('A::type finite_image => bool) => bool)
|
|
6967 |
(hollight.UNIV::'A::type finite_image => bool)" |
|
6968 |
by (import hollight FINITE_FINITE_IMAGE) |
|
6969 |
||
| 17652 | 6970 |
lemma DIMINDEX_NONZERO: "ALL s::'A::type => bool. dimindex s ~= 0" |
| 17644 | 6971 |
by (import hollight DIMINDEX_NONZERO) |
6972 |
||
| 17652 | 6973 |
lemma DIMINDEX_GE_1: "ALL x::'A::type => bool. <= (NUMERAL_BIT1 0) (dimindex x)" |
| 17644 | 6974 |
by (import hollight DIMINDEX_GE_1) |
6975 |
||
6976 |
lemma DIMINDEX_FINITE_IMAGE: "ALL (s::'A::type finite_image => bool) t::'A::type => bool. |
|
6977 |
dimindex s = dimindex t" |
|
6978 |
by (import hollight DIMINDEX_FINITE_IMAGE) |
|
6979 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
6980 |
definition finite_index :: "nat => 'A" where |
| 17644 | 6981 |
"(op ==::(nat => 'A::type) => (nat => 'A::type) => prop) |
6982 |
(finite_index::nat => 'A::type) |
|
6983 |
((Eps::((nat => 'A::type) => bool) => nat => 'A::type) |
|
6984 |
(%f::nat => 'A::type. |
|
6985 |
(All::('A::type => bool) => bool)
|
|
6986 |
(%x::'A::type. |
|
6987 |
(Ex1::(nat => bool) => bool) |
|
6988 |
(%n::nat. |
|
6989 |
(op &::bool => bool => bool) |
|
6990 |
((<=::nat => nat => bool) |
|
6991 |
((NUMERAL_BIT1::nat => nat) (0::nat)) n) |
|
6992 |
((op &::bool => bool => bool) |
|
6993 |
((<=::nat => nat => bool) n |
|
6994 |
((dimindex::('A::type => bool) => nat)
|
|
6995 |
(hollight.UNIV::'A::type => bool))) |
|
6996 |
((op =::'A::type => 'A::type => bool) (f n) x))))))" |
|
6997 |
||
6998 |
lemma DEF_finite_index: "(op =::(nat => 'A::type) => (nat => 'A::type) => bool) |
|
6999 |
(finite_index::nat => 'A::type) |
|
7000 |
((Eps::((nat => 'A::type) => bool) => nat => 'A::type) |
|
7001 |
(%f::nat => 'A::type. |
|
7002 |
(All::('A::type => bool) => bool)
|
|
7003 |
(%x::'A::type. |
|
7004 |
(Ex1::(nat => bool) => bool) |
|
7005 |
(%n::nat. |
|
7006 |
(op &::bool => bool => bool) |
|
7007 |
((<=::nat => nat => bool) |
|
7008 |
((NUMERAL_BIT1::nat => nat) (0::nat)) n) |
|
7009 |
((op &::bool => bool => bool) |
|
7010 |
((<=::nat => nat => bool) n |
|
7011 |
((dimindex::('A::type => bool) => nat)
|
|
7012 |
(hollight.UNIV::'A::type => bool))) |
|
7013 |
((op =::'A::type => 'A::type => bool) (f n) x))))))" |
|
7014 |
by (import hollight DEF_finite_index) |
|
7015 |
||
7016 |
lemma FINITE_INDEX_WORKS_FINITE: "(op -->::bool => bool => bool) |
|
7017 |
((FINITE::('N::type => bool) => bool) (hollight.UNIV::'N::type => bool))
|
|
7018 |
((All::('N::type => bool) => bool)
|
|
7019 |
(%x::'N::type. |
|
7020 |
(Ex1::(nat => bool) => bool) |
|
7021 |
(%xa::nat. |
|
7022 |
(op &::bool => bool => bool) |
|
7023 |
((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) |
|
7024 |
xa) |
|
7025 |
((op &::bool => bool => bool) |
|
7026 |
((<=::nat => nat => bool) xa |
|
7027 |
((dimindex::('N::type => bool) => nat)
|
|
7028 |
(hollight.UNIV::'N::type => bool))) |
|
7029 |
((op =::'N::type => 'N::type => bool) |
|
7030 |
((finite_index::nat => 'N::type) xa) x)))))" |
|
7031 |
by (import hollight FINITE_INDEX_WORKS_FINITE) |
|
7032 |
||
7033 |
lemma FINITE_INDEX_WORKS: "(All::('A::type finite_image => bool) => bool)
|
|
7034 |
(%i::'A::type finite_image. |
|
7035 |
(Ex1::(nat => bool) => bool) |
|
7036 |
(%n::nat. |
|
7037 |
(op &::bool => bool => bool) |
|
7038 |
((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) |
|
7039 |
n) |
|
7040 |
((op &::bool => bool => bool) |
|
7041 |
((<=::nat => nat => bool) n |
|
7042 |
((dimindex::('A::type => bool) => nat)
|
|
7043 |
(hollight.UNIV::'A::type => bool))) |
|
7044 |
((op =::'A::type finite_image => 'A::type finite_image => bool) |
|
7045 |
((finite_index::nat => 'A::type finite_image) n) i))))" |
|
7046 |
by (import hollight FINITE_INDEX_WORKS) |
|
7047 |
||
7048 |
lemma FINITE_INDEX_INJ: "(All::(nat => bool) => bool) |
|
7049 |
(%x::nat. |
|
7050 |
(All::(nat => bool) => bool) |
|
7051 |
(%xa::nat. |
|
7052 |
(op -->::bool => bool => bool) |
|
7053 |
((op &::bool => bool => bool) |
|
7054 |
((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) |
|
7055 |
x) |
|
7056 |
((op &::bool => bool => bool) |
|
7057 |
((<=::nat => nat => bool) x |
|
7058 |
((dimindex::('A::type => bool) => nat)
|
|
7059 |
(hollight.UNIV::'A::type => bool))) |
|
7060 |
((op &::bool => bool => bool) |
|
7061 |
((<=::nat => nat => bool) |
|
7062 |
((NUMERAL_BIT1::nat => nat) (0::nat)) xa) |
|
7063 |
((<=::nat => nat => bool) xa |
|
7064 |
((dimindex::('A::type => bool) => nat)
|
|
7065 |
(hollight.UNIV::'A::type => bool)))))) |
|
7066 |
((op =::bool => bool => bool) |
|
7067 |
((op =::'A::type => 'A::type => bool) |
|
7068 |
((finite_index::nat => 'A::type) x) |
|
7069 |
((finite_index::nat => 'A::type) xa)) |
|
7070 |
((op =::nat => nat => bool) x xa))))" |
|
7071 |
by (import hollight FINITE_INDEX_INJ) |
|
7072 |
||
7073 |
lemma FORALL_FINITE_INDEX: "(op =::bool => bool => bool) |
|
7074 |
((All::('N::type finite_image => bool) => bool)
|
|
7075 |
(P::'N::type finite_image => bool)) |
|
7076 |
((All::(nat => bool) => bool) |
|
7077 |
(%i::nat. |
|
7078 |
(op -->::bool => bool => bool) |
|
7079 |
((op &::bool => bool => bool) |
|
7080 |
((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) i) |
|
7081 |
((<=::nat => nat => bool) i |
|
7082 |
((dimindex::('N::type => bool) => nat)
|
|
7083 |
(hollight.UNIV::'N::type => bool)))) |
|
7084 |
(P ((finite_index::nat => 'N::type finite_image) i))))" |
|
7085 |
by (import hollight FORALL_FINITE_INDEX) |
|
7086 |
||
7087 |
typedef (open) ('A, 'B) cart = "(Collect::(('B::type finite_image => 'A::type) => bool)
|
|
7088 |
=> ('B::type finite_image => 'A::type) set)
|
|
7089 |
(%f::'B::type finite_image => 'A::type. True::bool)" morphisms "dest_cart" "mk_cart" |
|
7090 |
apply (rule light_ex_imp_nonempty[where t="(Eps::(('B::type finite_image => 'A::type) => bool)
|
|
7091 |
=> 'B::type finite_image => 'A::type) |
|
7092 |
(%f::'B::type finite_image => 'A::type. True::bool)"]) |
|
7093 |
by (import hollight TYDEF_cart) |
|
7094 |
||
7095 |
syntax |
|
7096 |
dest_cart :: _ |
|
7097 |
||
7098 |
syntax |
|
7099 |
mk_cart :: _ |
|
7100 |
||
7101 |
lemmas "TYDEF_cart_@intern" = typedef_hol2hollight |
|
| 17652 | 7102 |
[where a="a :: ('A, 'B) cart" and r=r ,
|
| 17644 | 7103 |
OF type_definition_cart] |
7104 |
||
7105 |
consts |
|
| 19093 | 7106 |
"$" :: "('q_46627, 'q_46634) cart => nat => 'q_46627" ("$")
|
| 17644 | 7107 |
|
7108 |
defs |
|
7109 |
"$_def": "$ == |
|
| 19093 | 7110 |
%(u::('q_46627::type, 'q_46634::type) cart) ua::nat.
|
| 17644 | 7111 |
dest_cart u (finite_index ua)" |
7112 |
||
7113 |
lemma "DEF_$": "$ = |
|
| 19093 | 7114 |
(%(u::('q_46627::type, 'q_46634::type) cart) ua::nat.
|
| 17644 | 7115 |
dest_cart u (finite_index ua))" |
7116 |
by (import hollight "DEF_$") |
|
7117 |
||
7118 |
lemma CART_EQ: "(All::(('A::type, 'B::type) cart => bool) => bool)
|
|
7119 |
(%x::('A::type, 'B::type) cart.
|
|
7120 |
(All::(('A::type, 'B::type) cart => bool) => bool)
|
|
7121 |
(%y::('A::type, 'B::type) cart.
|
|
7122 |
(op =::bool => bool => bool) |
|
7123 |
((op =::('A::type, 'B::type) cart
|
|
7124 |
=> ('A::type, 'B::type) cart => bool)
|
|
7125 |
x y) |
|
7126 |
((All::(nat => bool) => bool) |
|
7127 |
(%xa::nat. |
|
7128 |
(op -->::bool => bool => bool) |
|
7129 |
((op &::bool => bool => bool) |
|
7130 |
((<=::nat => nat => bool) |
|
7131 |
((NUMERAL_BIT1::nat => nat) (0::nat)) xa) |
|
7132 |
((<=::nat => nat => bool) xa |
|
7133 |
((dimindex::('B::type => bool) => nat)
|
|
7134 |
(hollight.UNIV::'B::type => bool)))) |
|
7135 |
((op =::'A::type => 'A::type => bool) |
|
7136 |
(($::('A::type, 'B::type) cart => nat => 'A::type) x xa)
|
|
7137 |
(($::('A::type, 'B::type) cart => nat => 'A::type) y
|
|
7138 |
xa))))))" |
|
7139 |
by (import hollight CART_EQ) |
|
7140 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7141 |
definition lambda :: "(nat => 'A) => ('A, 'B) cart" where
|
| 17644 | 7142 |
"(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart)
|
7143 |
=> ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop)
|
|
7144 |
(lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
|
|
7145 |
(%u::nat => 'A::type. |
|
7146 |
(Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
|
|
7147 |
(%f::('A::type, 'B::type) cart.
|
|
7148 |
(All::(nat => bool) => bool) |
|
7149 |
(%i::nat. |
|
7150 |
(op -->::bool => bool => bool) |
|
7151 |
((op &::bool => bool => bool) |
|
7152 |
((<=::nat => nat => bool) |
|
7153 |
((NUMERAL_BIT1::nat => nat) (0::nat)) i) |
|
7154 |
((<=::nat => nat => bool) i |
|
7155 |
((dimindex::('B::type => bool) => nat)
|
|
7156 |
(hollight.UNIV::'B::type => bool)))) |
|
7157 |
((op =::'A::type => 'A::type => bool) |
|
7158 |
(($::('A::type, 'B::type) cart => nat => 'A::type) f i)
|
|
7159 |
(u i)))))" |
|
7160 |
||
7161 |
lemma DEF_lambda: "(op =::((nat => 'A::type) => ('A::type, 'B::type) cart)
|
|
7162 |
=> ((nat => 'A::type) => ('A::type, 'B::type) cart) => bool)
|
|
7163 |
(lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
|
|
7164 |
(%u::nat => 'A::type. |
|
7165 |
(Eps::(('A::type, 'B::type) cart => bool) => ('A::type, 'B::type) cart)
|
|
7166 |
(%f::('A::type, 'B::type) cart.
|
|
7167 |
(All::(nat => bool) => bool) |
|
7168 |
(%i::nat. |
|
7169 |
(op -->::bool => bool => bool) |
|
7170 |
((op &::bool => bool => bool) |
|
7171 |
((<=::nat => nat => bool) |
|
7172 |
((NUMERAL_BIT1::nat => nat) (0::nat)) i) |
|
7173 |
((<=::nat => nat => bool) i |
|
7174 |
((dimindex::('B::type => bool) => nat)
|
|
7175 |
(hollight.UNIV::'B::type => bool)))) |
|
7176 |
((op =::'A::type => 'A::type => bool) |
|
7177 |
(($::('A::type, 'B::type) cart => nat => 'A::type) f i)
|
|
7178 |
(u i)))))" |
|
7179 |
by (import hollight DEF_lambda) |
|
7180 |
||
7181 |
lemma LAMBDA_BETA: "(All::(nat => bool) => bool) |
|
7182 |
(%x::nat. |
|
7183 |
(op -->::bool => bool => bool) |
|
7184 |
((op &::bool => bool => bool) |
|
7185 |
((<=::nat => nat => bool) ((NUMERAL_BIT1::nat => nat) (0::nat)) x) |
|
7186 |
((<=::nat => nat => bool) x |
|
7187 |
((dimindex::('B::type => bool) => nat)
|
|
7188 |
(hollight.UNIV::'B::type => bool)))) |
|
7189 |
((op =::'A::type => 'A::type => bool) |
|
7190 |
(($::('A::type, 'B::type) cart => nat => 'A::type)
|
|
7191 |
((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart)
|
|
7192 |
(g::nat => 'A::type)) |
|
7193 |
x) |
|
7194 |
(g x)))" |
|
7195 |
by (import hollight LAMBDA_BETA) |
|
7196 |
||
7197 |
lemma LAMBDA_UNIQUE: "(All::(('A::type, 'B::type) cart => bool) => bool)
|
|
7198 |
(%x::('A::type, 'B::type) cart.
|
|
7199 |
(All::((nat => 'A::type) => bool) => bool) |
|
7200 |
(%xa::nat => 'A::type. |
|
7201 |
(op =::bool => bool => bool) |
|
7202 |
((All::(nat => bool) => bool) |
|
7203 |
(%i::nat. |
|
7204 |
(op -->::bool => bool => bool) |
|
7205 |
((op &::bool => bool => bool) |
|
7206 |
((<=::nat => nat => bool) |
|
7207 |
((NUMERAL_BIT1::nat => nat) (0::nat)) i) |
|
7208 |
((<=::nat => nat => bool) i |
|
7209 |
((dimindex::('B::type => bool) => nat)
|
|
7210 |
(hollight.UNIV::'B::type => bool)))) |
|
7211 |
((op =::'A::type => 'A::type => bool) |
|
7212 |
(($::('A::type, 'B::type) cart => nat => 'A::type) x i)
|
|
7213 |
(xa i)))) |
|
7214 |
((op =::('A::type, 'B::type) cart
|
|
7215 |
=> ('A::type, 'B::type) cart => bool)
|
|
7216 |
((lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) xa)
|
|
7217 |
x)))" |
|
7218 |
by (import hollight LAMBDA_UNIQUE) |
|
7219 |
||
| 19093 | 7220 |
lemma LAMBDA_ETA: "ALL x::('q_46825::type, 'q_46829::type) cart. lambda ($ x) = x"
|
| 17644 | 7221 |
by (import hollight LAMBDA_ETA) |
7222 |
||
| 19093 | 7223 |
typedef (open) ('A, 'B) finite_sum = "(Collect::('A::type finite_image + 'B::type finite_image => bool)
|
7224 |
=> ('A::type finite_image + 'B::type finite_image) set)
|
|
7225 |
(%x::'A::type finite_image + 'B::type finite_image. True::bool)" morphisms "dest_finite_sum" "mk_finite_sum" |
|
7226 |
apply (rule light_ex_imp_nonempty[where t="(Eps::('A::type finite_image + 'B::type finite_image => bool)
|
|
7227 |
=> 'A::type finite_image + 'B::type finite_image) |
|
7228 |
(%x::'A::type finite_image + 'B::type finite_image. True::bool)"]) |
|
| 17644 | 7229 |
by (import hollight TYDEF_finite_sum) |
7230 |
||
7231 |
syntax |
|
7232 |
dest_finite_sum :: _ |
|
7233 |
||
7234 |
syntax |
|
7235 |
mk_finite_sum :: _ |
|
7236 |
||
7237 |
lemmas "TYDEF_finite_sum_@intern" = typedef_hol2hollight |
|
| 17652 | 7238 |
[where a="a :: ('A, 'B) finite_sum" and r=r ,
|
| 17644 | 7239 |
OF type_definition_finite_sum] |
7240 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7241 |
definition pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where
|
| 17644 | 7242 |
"(op ==::(('A::type, 'M::type) cart
|
7243 |
=> ('A::type, 'N::type) cart
|
|
7244 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7245 |
=> (('A::type, 'M::type) cart
|
|
7246 |
=> ('A::type, 'N::type) cart
|
|
7247 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7248 |
=> prop) |
|
7249 |
(pastecart::('A::type, 'M::type) cart
|
|
7250 |
=> ('A::type, 'N::type) cart
|
|
7251 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7252 |
(%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
|
|
7253 |
(lambda::(nat => 'A::type) |
|
7254 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7255 |
(%i::nat. |
|
7256 |
(COND::bool => 'A::type => 'A::type => 'A::type) |
|
7257 |
((<=::nat => nat => bool) i |
|
7258 |
((dimindex::('M::type => bool) => nat)
|
|
7259 |
(hollight.UNIV::'M::type => bool))) |
|
7260 |
(($::('A::type, 'M::type) cart => nat => 'A::type) u i)
|
|
7261 |
(($::('A::type, 'N::type) cart => nat => 'A::type) ua
|
|
7262 |
((op -::nat => nat => nat) i |
|
7263 |
((dimindex::('M::type => bool) => nat)
|
|
7264 |
(hollight.UNIV::'M::type => bool))))))" |
|
7265 |
||
7266 |
lemma DEF_pastecart: "(op =::(('A::type, 'M::type) cart
|
|
7267 |
=> ('A::type, 'N::type) cart
|
|
7268 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7269 |
=> (('A::type, 'M::type) cart
|
|
7270 |
=> ('A::type, 'N::type) cart
|
|
7271 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7272 |
=> bool) |
|
7273 |
(pastecart::('A::type, 'M::type) cart
|
|
7274 |
=> ('A::type, 'N::type) cart
|
|
7275 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7276 |
(%(u::('A::type, 'M::type) cart) ua::('A::type, 'N::type) cart.
|
|
7277 |
(lambda::(nat => 'A::type) |
|
7278 |
=> ('A::type, ('M::type, 'N::type) finite_sum) cart)
|
|
7279 |
(%i::nat. |
|
7280 |
(COND::bool => 'A::type => 'A::type => 'A::type) |
|
7281 |
((<=::nat => nat => bool) i |
|
7282 |
((dimindex::('M::type => bool) => nat)
|
|
7283 |
(hollight.UNIV::'M::type => bool))) |
|
7284 |
(($::('A::type, 'M::type) cart => nat => 'A::type) u i)
|
|
7285 |
(($::('A::type, 'N::type) cart => nat => 'A::type) ua
|
|
7286 |
((op -::nat => nat => nat) i |
|
7287 |
((dimindex::('M::type => bool) => nat)
|
|
7288 |
(hollight.UNIV::'M::type => bool))))))" |
|
7289 |
by (import hollight DEF_pastecart) |
|
7290 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7291 |
definition fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where
|
| 17644 | 7292 |
"fstcart == |
7293 |
%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u)"
|
|
7294 |
||
7295 |
lemma DEF_fstcart: "fstcart = |
|
7296 |
(%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u))"
|
|
7297 |
by (import hollight DEF_fstcart) |
|
7298 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7299 |
definition sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where
|
| 17644 | 7300 |
"(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart
|
7301 |
=> ('A::type, 'N::type) cart)
|
|
7302 |
=> (('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7303 |
=> ('A::type, 'N::type) cart)
|
|
7304 |
=> prop) |
|
7305 |
(sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7306 |
=> ('A::type, 'N::type) cart)
|
|
7307 |
(%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
|
|
7308 |
(lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
|
|
7309 |
(%i::nat. |
|
7310 |
($::('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7311 |
=> nat => 'A::type) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7312 |
u ((plus::nat => nat => nat) i |
| 17644 | 7313 |
((dimindex::('M::type => bool) => nat)
|
7314 |
(hollight.UNIV::'M::type => bool)))))" |
|
7315 |
||
7316 |
lemma DEF_sndcart: "(op =::(('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7317 |
=> ('A::type, 'N::type) cart)
|
|
7318 |
=> (('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7319 |
=> ('A::type, 'N::type) cart)
|
|
7320 |
=> bool) |
|
7321 |
(sndcart::('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7322 |
=> ('A::type, 'N::type) cart)
|
|
7323 |
(%u::('A::type, ('M::type, 'N::type) finite_sum) cart.
|
|
7324 |
(lambda::(nat => 'A::type) => ('A::type, 'N::type) cart)
|
|
7325 |
(%i::nat. |
|
7326 |
($::('A::type, ('M::type, 'N::type) finite_sum) cart
|
|
7327 |
=> nat => 'A::type) |
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7328 |
u ((plus::nat => nat => nat) i |
| 17644 | 7329 |
((dimindex::('M::type => bool) => nat)
|
7330 |
(hollight.UNIV::'M::type => bool)))))" |
|
7331 |
by (import hollight DEF_sndcart) |
|
7332 |
||
7333 |
lemma DIMINDEX_HAS_SIZE_FINITE_SUM: "(HAS_SIZE::(('M::type, 'N::type) finite_sum => bool) => nat => bool)
|
|
7334 |
(hollight.UNIV::('M::type, 'N::type) finite_sum => bool)
|
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7335 |
((plus::nat => nat => nat) |
| 17644 | 7336 |
((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool))
|
7337 |
((dimindex::('N::type => bool) => nat)
|
|
7338 |
(hollight.UNIV::'N::type => bool)))" |
|
7339 |
by (import hollight DIMINDEX_HAS_SIZE_FINITE_SUM) |
|
7340 |
||
7341 |
lemma DIMINDEX_FINITE_SUM: "(op =::nat => nat => bool) |
|
7342 |
((dimindex::(('M::type, 'N::type) finite_sum => bool) => nat)
|
|
7343 |
(hollight.UNIV::('M::type, 'N::type) finite_sum => bool))
|
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7344 |
((plus::nat => nat => nat) |
| 17644 | 7345 |
((dimindex::('M::type => bool) => nat) (hollight.UNIV::'M::type => bool))
|
7346 |
((dimindex::('N::type => bool) => nat)
|
|
7347 |
(hollight.UNIV::'N::type => bool)))" |
|
7348 |
by (import hollight DIMINDEX_FINITE_SUM) |
|
7349 |
||
7350 |
lemma FSTCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart.
|
|
7351 |
fstcart (pastecart x xa) = x" |
|
7352 |
by (import hollight FSTCART_PASTECART) |
|
7353 |
||
7354 |
lemma SNDCART_PASTECART: "ALL (x::('A::type, 'M::type) cart) xa::('A::type, 'N::type) cart.
|
|
7355 |
sndcart (pastecart x xa) = xa" |
|
7356 |
by (import hollight SNDCART_PASTECART) |
|
7357 |
||
| 19093 | 7358 |
lemma PASTECART_FST_SND: "ALL x::('q_47149::type, ('q_47146::type, 'q_47144::type) finite_sum) cart.
|
| 17644 | 7359 |
pastecart (fstcart x) (sndcart x) = x" |
7360 |
by (import hollight PASTECART_FST_SND) |
|
7361 |
||
| 19093 | 7362 |
lemma PASTECART_EQ: "ALL (x::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart)
|
7363 |
y::('q_47187::type, ('q_47177::type, 'q_47188::type) finite_sum) cart.
|
|
| 17644 | 7364 |
(x = y) = (fstcart x = fstcart y & sndcart x = sndcart y)" |
7365 |
by (import hollight PASTECART_EQ) |
|
7366 |
||
| 19093 | 7367 |
lemma FORALL_PASTECART: "All (P::('q_47208::type, ('q_47209::type, 'q_47210::type) finite_sum) cart
|
| 17644 | 7368 |
=> bool) = |
| 19093 | 7369 |
(ALL (x::('q_47208::type, 'q_47209::type) cart)
|
7370 |
y::('q_47208::type, 'q_47210::type) cart. P (pastecart x y))"
|
|
| 17644 | 7371 |
by (import hollight FORALL_PASTECART) |
7372 |
||
| 19093 | 7373 |
lemma EXISTS_PASTECART: "Ex (P::('q_47230::type, ('q_47231::type, 'q_47232::type) finite_sum) cart
|
| 17644 | 7374 |
=> bool) = |
| 19093 | 7375 |
(EX (x::('q_47230::type, 'q_47231::type) cart)
|
7376 |
y::('q_47230::type, 'q_47232::type) cart. P (pastecart x y))"
|
|
| 17644 | 7377 |
by (import hollight EXISTS_PASTECART) |
7378 |
||
7379 |
lemma SURJECTIVE_IFF_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type. |
|
7380 |
FINITE s & FINITE t & CARD s = CARD t & SUBSET (IMAGE f s) t --> |
|
7381 |
(ALL y::'B::type. IN y t --> (EX x::'A::type. IN x s & f x = y)) = |
|
7382 |
(ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y)" |
|
7383 |
by (import hollight SURJECTIVE_IFF_INJECTIVE_GEN) |
|
7384 |
||
7385 |
lemma SURJECTIVE_IFF_INJECTIVE: "ALL (x::'A::type => bool) xa::'A::type => 'A::type. |
|
7386 |
FINITE x & SUBSET (IMAGE xa x) x --> |
|
7387 |
(ALL y::'A::type. IN y x --> (EX xb::'A::type. IN xb x & xa xb = y)) = |
|
7388 |
(ALL (xb::'A::type) y::'A::type. |
|
7389 |
IN xb x & IN y x & xa xb = xa y --> xb = y)" |
|
7390 |
by (import hollight SURJECTIVE_IFF_INJECTIVE) |
|
7391 |
||
7392 |
lemma IMAGE_IMP_INJECTIVE_GEN: "ALL (s::'A::type => bool) (t::'B::type => bool) f::'A::type => 'B::type. |
|
7393 |
FINITE s & CARD s = CARD t & IMAGE f s = t --> |
|
7394 |
(ALL (x::'A::type) y::'A::type. IN x s & IN y s & f x = f y --> x = y)" |
|
7395 |
by (import hollight IMAGE_IMP_INJECTIVE_GEN) |
|
7396 |
||
| 19093 | 7397 |
lemma IMAGE_IMP_INJECTIVE: "ALL (s::'q_47557::type => bool) f::'q_47557::type => 'q_47557::type. |
| 17644 | 7398 |
FINITE s & IMAGE f s = s --> |
| 19093 | 7399 |
(ALL (x::'q_47557::type) y::'q_47557::type. |
| 17644 | 7400 |
IN x s & IN y s & f x = f y --> x = y)" |
7401 |
by (import hollight IMAGE_IMP_INJECTIVE) |
|
7402 |
||
7403 |
lemma CARD_LE_INJ: "ALL (x::'A::type => bool) xa::'B::type => bool. |
|
7404 |
FINITE x & FINITE xa & <= (CARD x) (CARD xa) --> |
|
7405 |
(EX f::'A::type => 'B::type. |
|
7406 |
SUBSET (IMAGE f x) xa & |
|
7407 |
(ALL (xa::'A::type) y::'A::type. |
|
7408 |
IN xa x & IN y x & f xa = f y --> xa = y))" |
|
7409 |
by (import hollight CARD_LE_INJ) |
|
7410 |
||
| 19093 | 7411 |
lemma FORALL_IN_CLAUSES: "(ALL x::'q_47663::type => bool. |
7412 |
(ALL xa::'q_47663::type. IN xa EMPTY --> x xa) = True) & |
|
7413 |
(ALL (x::'q_47703::type => bool) (xa::'q_47703::type) |
|
7414 |
xb::'q_47703::type => bool. |
|
7415 |
(ALL xc::'q_47703::type. IN xc (INSERT xa xb) --> x xc) = |
|
7416 |
(x xa & (ALL xa::'q_47703::type. IN xa xb --> x xa)))" |
|
| 17644 | 7417 |
by (import hollight FORALL_IN_CLAUSES) |
7418 |
||
| 19093 | 7419 |
lemma EXISTS_IN_CLAUSES: "(ALL x::'q_47723::type => bool. |
7420 |
(EX xa::'q_47723::type. IN xa EMPTY & x xa) = False) & |
|
7421 |
(ALL (x::'q_47763::type => bool) (xa::'q_47763::type) |
|
7422 |
xb::'q_47763::type => bool. |
|
7423 |
(EX xc::'q_47763::type. IN xc (INSERT xa xb) & x xc) = |
|
7424 |
(x xa | (EX xa::'q_47763::type. IN xa xb & x xa)))" |
|
| 17644 | 7425 |
by (import hollight EXISTS_IN_CLAUSES) |
7426 |
||
| 19093 | 7427 |
lemma SURJECTIVE_ON_RIGHT_INVERSE: "ALL (x::'q_47819::type => 'q_47820::type) xa::'q_47820::type => bool. |
7428 |
(ALL xb::'q_47820::type. |
|
| 17644 | 7429 |
IN xb xa --> |
| 19093 | 7430 |
(EX xa::'q_47819::type. |
7431 |
IN xa (s::'q_47819::type => bool) & x xa = xb)) = |
|
7432 |
(EX g::'q_47820::type => 'q_47819::type. |
|
7433 |
ALL y::'q_47820::type. IN y xa --> IN (g y) s & x (g y) = y)" |
|
| 17644 | 7434 |
by (import hollight SURJECTIVE_ON_RIGHT_INVERSE) |
7435 |
||
| 19093 | 7436 |
lemma INJECTIVE_ON_LEFT_INVERSE: "ALL (x::'q_47913::type => 'q_47916::type) xa::'q_47913::type => bool. |
7437 |
(ALL (xb::'q_47913::type) y::'q_47913::type. |
|
| 17644 | 7438 |
IN xb xa & IN y xa & x xb = x y --> xb = y) = |
| 19093 | 7439 |
(EX xb::'q_47916::type => 'q_47913::type. |
7440 |
ALL xc::'q_47913::type. IN xc xa --> xb (x xc) = xc)" |
|
| 17644 | 7441 |
by (import hollight INJECTIVE_ON_LEFT_INVERSE) |
7442 |
||
| 19093 | 7443 |
lemma SURJECTIVE_RIGHT_INVERSE: "(ALL y::'q_47941::type. |
7444 |
EX x::'q_47944::type. (f::'q_47944::type => 'q_47941::type) x = y) = |
|
7445 |
(EX g::'q_47941::type => 'q_47944::type. ALL y::'q_47941::type. f (g y) = y)" |
|
| 17644 | 7446 |
by (import hollight SURJECTIVE_RIGHT_INVERSE) |
7447 |
||
| 19093 | 7448 |
lemma INJECTIVE_LEFT_INVERSE: "(ALL (x::'q_47978::type) xa::'q_47978::type. |
7449 |
(f::'q_47978::type => 'q_47981::type) x = f xa --> x = xa) = |
|
7450 |
(EX g::'q_47981::type => 'q_47978::type. ALL x::'q_47978::type. g (f x) = x)" |
|
| 17644 | 7451 |
by (import hollight INJECTIVE_LEFT_INVERSE) |
7452 |
||
| 19093 | 7453 |
lemma FUNCTION_FACTORS_RIGHT: "ALL (x::'q_48017::type => 'q_48018::type) |
7454 |
xa::'q_48005::type => 'q_48018::type. |
|
7455 |
(ALL xb::'q_48017::type. EX y::'q_48005::type. xa y = x xb) = |
|
7456 |
(EX xb::'q_48017::type => 'q_48005::type. x = xa o xb)" |
|
| 17644 | 7457 |
by (import hollight FUNCTION_FACTORS_RIGHT) |
7458 |
||
| 19093 | 7459 |
lemma FUNCTION_FACTORS_LEFT: "ALL (x::'q_48090::type => 'q_48091::type) |
7460 |
xa::'q_48090::type => 'q_48070::type. |
|
7461 |
(ALL (xb::'q_48090::type) y::'q_48090::type. |
|
| 17644 | 7462 |
xa xb = xa y --> x xb = x y) = |
| 19093 | 7463 |
(EX xb::'q_48070::type => 'q_48091::type. x = xb o xa)" |
| 17644 | 7464 |
by (import hollight FUNCTION_FACTORS_LEFT) |
7465 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7466 |
definition dotdot :: "nat => nat => nat => bool" where |
| 17644 | 7467 |
"dotdot == |
7468 |
%(u::nat) ua::nat. |
|
7469 |
GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x)" |
|
7470 |
||
7471 |
lemma DEF__dot__dot_: "dotdot = |
|
7472 |
(%(u::nat) ua::nat. |
|
7473 |
GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x))" |
|
7474 |
by (import hollight DEF__dot__dot_) |
|
7475 |
||
7476 |
lemma FINITE_NUMSEG: "ALL (m::nat) n::nat. FINITE (dotdot m n)" |
|
7477 |
by (import hollight FINITE_NUMSEG) |
|
7478 |
||
| 19093 | 7479 |
lemma NUMSEG_COMBINE_R: "ALL (x::'q_48166::type) (p::nat) m::nat. |
| 17644 | 7480 |
<= m p & <= p (n::nat) --> |
| 17652 | 7481 |
hollight.UNION (dotdot m p) (dotdot (p + NUMERAL_BIT1 0) n) = dotdot m n" |
| 17644 | 7482 |
by (import hollight NUMSEG_COMBINE_R) |
7483 |
||
| 19093 | 7484 |
lemma NUMSEG_COMBINE_L: "ALL (x::'q_48204::type) (p::nat) m::nat. |
| 17644 | 7485 |
<= m p & <= p (n::nat) --> |
| 17652 | 7486 |
hollight.UNION (dotdot m (p - NUMERAL_BIT1 0)) (dotdot p n) = dotdot m n" |
| 17644 | 7487 |
by (import hollight NUMSEG_COMBINE_L) |
7488 |
||
7489 |
lemma NUMSEG_LREC: "ALL (x::nat) xa::nat. |
|
| 17652 | 7490 |
<= x xa --> INSERT x (dotdot (x + NUMERAL_BIT1 0) xa) = dotdot x xa" |
| 17644 | 7491 |
by (import hollight NUMSEG_LREC) |
7492 |
||
7493 |
lemma NUMSEG_RREC: "ALL (x::nat) xa::nat. |
|
| 17652 | 7494 |
<= x xa --> INSERT xa (dotdot x (xa - NUMERAL_BIT1 0)) = dotdot x xa" |
| 17644 | 7495 |
by (import hollight NUMSEG_RREC) |
7496 |
||
7497 |
lemma NUMSEG_REC: "ALL (x::nat) xa::nat. |
|
7498 |
<= x (Suc xa) --> dotdot x (Suc xa) = INSERT (Suc xa) (dotdot x xa)" |
|
7499 |
by (import hollight NUMSEG_REC) |
|
7500 |
||
7501 |
lemma IN_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat. IN xb (dotdot x xa) = (<= x xb & <= xb xa)" |
|
7502 |
by (import hollight IN_NUMSEG) |
|
7503 |
||
7504 |
lemma NUMSEG_SING: "ALL x::nat. dotdot x x = INSERT x EMPTY" |
|
7505 |
by (import hollight NUMSEG_SING) |
|
7506 |
||
7507 |
lemma NUMSEG_EMPTY: "ALL (x::nat) xa::nat. (dotdot x xa = EMPTY) = < xa x" |
|
7508 |
by (import hollight NUMSEG_EMPTY) |
|
7509 |
||
| 17652 | 7510 |
lemma CARD_NUMSEG_LEMMA: "ALL (m::nat) d::nat. CARD (dotdot m (m + d)) = d + NUMERAL_BIT1 0" |
| 17644 | 7511 |
by (import hollight CARD_NUMSEG_LEMMA) |
7512 |
||
| 17652 | 7513 |
lemma CARD_NUMSEG: "ALL (m::nat) n::nat. CARD (dotdot m n) = n + NUMERAL_BIT1 0 - m" |
| 17644 | 7514 |
by (import hollight CARD_NUMSEG) |
7515 |
||
| 17652 | 7516 |
lemma HAS_SIZE_NUMSEG: "ALL (x::nat) xa::nat. HAS_SIZE (dotdot x xa) (xa + NUMERAL_BIT1 0 - x)" |
| 17644 | 7517 |
by (import hollight HAS_SIZE_NUMSEG) |
7518 |
||
| 17652 | 7519 |
lemma CARD_NUMSEG_1: "ALL x::nat. CARD (dotdot (NUMERAL_BIT1 0) x) = x" |
| 17644 | 7520 |
by (import hollight CARD_NUMSEG_1) |
7521 |
||
| 17652 | 7522 |
lemma HAS_SIZE_NUMSEG_1: "ALL x::nat. HAS_SIZE (dotdot (NUMERAL_BIT1 0) x) x" |
| 17644 | 7523 |
by (import hollight HAS_SIZE_NUMSEG_1) |
7524 |
||
| 17652 | 7525 |
lemma NUMSEG_CLAUSES: "(ALL m::nat. dotdot m 0 = COND (m = 0) (INSERT 0 EMPTY) EMPTY) & |
| 17644 | 7526 |
(ALL (m::nat) n::nat. |
7527 |
dotdot m (Suc n) = |
|
7528 |
COND (<= m (Suc n)) (INSERT (Suc n) (dotdot m n)) (dotdot m n))" |
|
7529 |
by (import hollight NUMSEG_CLAUSES) |
|
7530 |
||
7531 |
lemma FINITE_INDEX_NUMSEG: "ALL s::'A::type => bool. |
|
7532 |
FINITE s = |
|
7533 |
(EX f::nat => 'A::type. |
|
7534 |
(ALL (i::nat) j::nat. |
|
| 17652 | 7535 |
IN i (dotdot (NUMERAL_BIT1 0) (CARD s)) & |
7536 |
IN j (dotdot (NUMERAL_BIT1 0) (CARD s)) & f i = f j --> |
|
| 17644 | 7537 |
i = j) & |
| 17652 | 7538 |
s = IMAGE f (dotdot (NUMERAL_BIT1 0) (CARD s)))" |
| 17644 | 7539 |
by (import hollight FINITE_INDEX_NUMSEG) |
7540 |
||
7541 |
lemma FINITE_INDEX_NUMBERS: "ALL s::'A::type => bool. |
|
7542 |
FINITE s = |
|
7543 |
(EX (k::nat => bool) f::nat => 'A::type. |
|
7544 |
(ALL (i::nat) j::nat. IN i k & IN j k & f i = f j --> i = j) & |
|
7545 |
FINITE k & s = IMAGE f k)" |
|
7546 |
by (import hollight FINITE_INDEX_NUMBERS) |
|
7547 |
||
7548 |
lemma DISJOINT_NUMSEG: "ALL (x::nat) (xa::nat) (xb::nat) xc::nat. |
|
7549 |
DISJOINT (dotdot x xa) (dotdot xb xc) = |
|
7550 |
(< xa xb | < xc x | < xa x | < xc xb)" |
|
7551 |
by (import hollight DISJOINT_NUMSEG) |
|
7552 |
||
7553 |
lemma NUMSEG_ADD_SPLIT: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 7554 |
<= x (xa + NUMERAL_BIT1 0) --> |
| 17644 | 7555 |
dotdot x (xa + xb) = |
| 17652 | 7556 |
hollight.UNION (dotdot x xa) (dotdot (xa + NUMERAL_BIT1 0) (xa + xb))" |
| 17644 | 7557 |
by (import hollight NUMSEG_ADD_SPLIT) |
7558 |
||
7559 |
lemma NUMSEG_OFFSET_IMAGE: "ALL (x::nat) (xa::nat) xb::nat. |
|
7560 |
dotdot (x + xb) (xa + xb) = IMAGE (%i::nat. i + xb) (dotdot x xa)" |
|
7561 |
by (import hollight NUMSEG_OFFSET_IMAGE) |
|
7562 |
||
7563 |
lemma SUBSET_NUMSEG: "ALL (m::nat) (n::nat) (p::nat) q::nat. |
|
7564 |
SUBSET (dotdot m n) (dotdot p q) = (< n m | <= p m & <= n q)" |
|
7565 |
by (import hollight SUBSET_NUMSEG) |
|
7566 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7567 |
definition neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" where
|
| 17644 | 7568 |
"neutral == |
| 19093 | 7569 |
%u::'q_48985::type => 'q_48985::type => 'q_48985::type. |
7570 |
SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y" |
|
| 17644 | 7571 |
|
7572 |
lemma DEF_neutral: "neutral = |
|
| 19093 | 7573 |
(%u::'q_48985::type => 'q_48985::type => 'q_48985::type. |
7574 |
SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y)" |
|
| 17644 | 7575 |
by (import hollight DEF_neutral) |
7576 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7577 |
definition monoidal :: "('A => 'A => 'A) => bool" where
|
| 17644 | 7578 |
"monoidal == |
7579 |
%u::'A::type => 'A::type => 'A::type. |
|
7580 |
(ALL (x::'A::type) y::'A::type. u x y = u y x) & |
|
7581 |
(ALL (x::'A::type) (y::'A::type) z::'A::type. |
|
7582 |
u x (u y z) = u (u x y) z) & |
|
7583 |
(ALL x::'A::type. u (neutral u) x = x)" |
|
7584 |
||
7585 |
lemma DEF_monoidal: "monoidal = |
|
7586 |
(%u::'A::type => 'A::type => 'A::type. |
|
7587 |
(ALL (x::'A::type) y::'A::type. u x y = u y x) & |
|
7588 |
(ALL (x::'A::type) (y::'A::type) z::'A::type. |
|
7589 |
u x (u y z) = u (u x y) z) & |
|
7590 |
(ALL x::'A::type. u (neutral u) x = x))" |
|
7591 |
by (import hollight DEF_monoidal) |
|
7592 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7593 |
definition support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where
|
| 17644 | 7594 |
"support == |
7595 |
%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type) |
|
7596 |
ub::'A::type => bool. |
|
7597 |
GSPEC |
|
7598 |
(%uc::'A::type. |
|
7599 |
EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x)" |
|
7600 |
||
7601 |
lemma DEF_support: "support = |
|
7602 |
(%(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type) |
|
7603 |
ub::'A::type => bool. |
|
7604 |
GSPEC |
|
7605 |
(%uc::'A::type. |
|
7606 |
EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x))" |
|
7607 |
by (import hollight DEF_support) |
|
7608 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7609 |
definition iterate :: "('q_49090 => 'q_49090 => 'q_49090)
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7610 |
=> ('A => bool) => ('A => 'q_49090) => 'q_49090" where
|
| 17644 | 7611 |
"iterate == |
| 19093 | 7612 |
%(u::'q_49090::type => 'q_49090::type => 'q_49090::type) |
7613 |
(ua::'A::type => bool) ub::'A::type => 'q_49090::type. |
|
| 17644 | 7614 |
ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u)" |
7615 |
||
7616 |
lemma DEF_iterate: "iterate = |
|
| 19093 | 7617 |
(%(u::'q_49090::type => 'q_49090::type => 'q_49090::type) |
7618 |
(ua::'A::type => bool) ub::'A::type => 'q_49090::type. |
|
| 17644 | 7619 |
ITSET (%x::'A::type. u (ub x)) (support u ub ua) (neutral u))" |
7620 |
by (import hollight DEF_iterate) |
|
7621 |
||
| 19093 | 7622 |
lemma IN_SUPPORT: "ALL (x::'q_49133::type => 'q_49133::type => 'q_49133::type) |
7623 |
(xa::'q_49136::type => 'q_49133::type) (xb::'q_49136::type) |
|
7624 |
xc::'q_49136::type => bool. |
|
| 17644 | 7625 |
IN xb (support x xa xc) = (IN xb xc & xa xb ~= neutral x)" |
7626 |
by (import hollight IN_SUPPORT) |
|
7627 |
||
| 19093 | 7628 |
lemma SUPPORT_SUPPORT: "ALL (x::'q_49158::type => 'q_49158::type => 'q_49158::type) |
7629 |
(xa::'q_49169::type => 'q_49158::type) xb::'q_49169::type => bool. |
|
| 17644 | 7630 |
support x xa (support x xa xb) = support x xa xb" |
7631 |
by (import hollight SUPPORT_SUPPORT) |
|
7632 |
||
| 19093 | 7633 |
lemma SUPPORT_EMPTY: "ALL (x::'q_49194::type => 'q_49194::type => 'q_49194::type) |
7634 |
(xa::'q_49208::type => 'q_49194::type) xb::'q_49208::type => bool. |
|
7635 |
(ALL xc::'q_49208::type. IN xc xb --> xa xc = neutral x) = |
|
| 17644 | 7636 |
(support x xa xb = EMPTY)" |
7637 |
by (import hollight SUPPORT_EMPTY) |
|
7638 |
||
| 19093 | 7639 |
lemma SUPPORT_SUBSET: "ALL (x::'q_49228::type => 'q_49228::type => 'q_49228::type) |
7640 |
(xa::'q_49229::type => 'q_49228::type) xb::'q_49229::type => bool. |
|
| 17644 | 7641 |
SUBSET (support x xa xb) xb" |
7642 |
by (import hollight SUPPORT_SUBSET) |
|
7643 |
||
| 19093 | 7644 |
lemma FINITE_SUPPORT: "ALL (u::'q_49252::type => 'q_49252::type => 'q_49252::type) |
7645 |
(f::'q_49246::type => 'q_49252::type) s::'q_49246::type => bool. |
|
| 17644 | 7646 |
FINITE s --> FINITE (support u f s)" |
7647 |
by (import hollight FINITE_SUPPORT) |
|
7648 |
||
| 19093 | 7649 |
lemma SUPPORT_CLAUSES: "(ALL x::'q_49270::type => 'q_49300::type. |
7650 |
support (u_4247::'q_49300::type => 'q_49300::type => 'q_49300::type) x |
|
| 17644 | 7651 |
EMPTY = |
7652 |
EMPTY) & |
|
| 19093 | 7653 |
(ALL (x::'q_49318::type => 'q_49300::type) (xa::'q_49318::type) |
7654 |
xb::'q_49318::type => bool. |
|
7655 |
support u_4247 x (INSERT xa xb) = |
|
7656 |
COND (x xa = neutral u_4247) (support u_4247 x xb) |
|
7657 |
(INSERT xa (support u_4247 x xb))) & |
|
7658 |
(ALL (x::'q_49351::type => 'q_49300::type) (xa::'q_49351::type) |
|
7659 |
xb::'q_49351::type => bool. |
|
7660 |
support u_4247 x (DELETE xb xa) = DELETE (support u_4247 x xb) xa) & |
|
7661 |
(ALL (x::'q_49389::type => 'q_49300::type) (xa::'q_49389::type => bool) |
|
7662 |
xb::'q_49389::type => bool. |
|
7663 |
support u_4247 x (hollight.UNION xa xb) = |
|
7664 |
hollight.UNION (support u_4247 x xa) (support u_4247 x xb)) & |
|
7665 |
(ALL (x::'q_49427::type => 'q_49300::type) (xa::'q_49427::type => bool) |
|
7666 |
xb::'q_49427::type => bool. |
|
7667 |
support u_4247 x (hollight.INTER xa xb) = |
|
7668 |
hollight.INTER (support u_4247 x xa) (support u_4247 x xb)) & |
|
7669 |
(ALL (x::'q_49463::type => 'q_49300::type) (xa::'q_49463::type => bool) |
|
7670 |
xb::'q_49463::type => bool. |
|
7671 |
support u_4247 x (DIFF xa xb) = |
|
7672 |
DIFF (support u_4247 x xa) (support u_4247 x xb))" |
|
| 17644 | 7673 |
by (import hollight SUPPORT_CLAUSES) |
7674 |
||
| 19093 | 7675 |
lemma ITERATE_SUPPORT: "ALL (x::'q_49484::type => 'q_49484::type => 'q_49484::type) |
7676 |
(xa::'q_49485::type => 'q_49484::type) xb::'q_49485::type => bool. |
|
| 17644 | 7677 |
FINITE (support x xa xb) --> |
7678 |
iterate x (support x xa xb) xa = iterate x xb xa" |
|
7679 |
by (import hollight ITERATE_SUPPORT) |
|
7680 |
||
| 19093 | 7681 |
lemma SUPPORT_DELTA: "ALL (x::'q_49529::type => 'q_49529::type => 'q_49529::type) |
7682 |
(xa::'q_49557::type => bool) (xb::'q_49557::type => 'q_49529::type) |
|
7683 |
xc::'q_49557::type. |
|
7684 |
support x (%xa::'q_49557::type. COND (xa = xc) (xb xa) (neutral x)) xa = |
|
| 17644 | 7685 |
COND (IN xc xa) (support x xb (INSERT xc EMPTY)) EMPTY" |
7686 |
by (import hollight SUPPORT_DELTA) |
|
7687 |
||
| 19093 | 7688 |
lemma FINITE_SUPPORT_DELTA: "ALL (x::'q_49578::type => 'q_49578::type => 'q_49578::type) |
7689 |
(xa::'q_49587::type => 'q_49578::type) xb::'q_49587::type. |
|
| 17644 | 7690 |
FINITE |
| 19093 | 7691 |
(support x (%xc::'q_49587::type. COND (xc = xb) (xa xc) (neutral x)) |
7692 |
(s::'q_49587::type => bool))" |
|
| 17644 | 7693 |
by (import hollight FINITE_SUPPORT_DELTA) |
7694 |
||
| 19093 | 7695 |
lemma ITERATE_CLAUSES_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7696 |
monoidal u_4247 --> |
|
7697 |
(ALL f::'A::type => 'B::type. iterate u_4247 EMPTY f = neutral u_4247) & |
|
| 17644 | 7698 |
(ALL (f::'A::type => 'B::type) (x::'A::type) s::'A::type => bool. |
| 19093 | 7699 |
monoidal u_4247 & FINITE (support u_4247 f s) --> |
7700 |
iterate u_4247 (INSERT x s) f = |
|
7701 |
COND (IN x s) (iterate u_4247 s f) |
|
7702 |
(u_4247 (f x) (iterate u_4247 s f)))" |
|
| 17644 | 7703 |
by (import hollight ITERATE_CLAUSES_GEN) |
7704 |
||
| 19093 | 7705 |
lemma ITERATE_CLAUSES: "ALL x::'q_49755::type => 'q_49755::type => 'q_49755::type. |
| 17644 | 7706 |
monoidal x --> |
| 19093 | 7707 |
(ALL f::'q_49713::type => 'q_49755::type. |
| 17644 | 7708 |
iterate x EMPTY f = neutral x) & |
| 19093 | 7709 |
(ALL (f::'q_49757::type => 'q_49755::type) (xa::'q_49757::type) |
7710 |
s::'q_49757::type => bool. |
|
| 17644 | 7711 |
FINITE s --> |
7712 |
iterate x (INSERT xa s) f = |
|
7713 |
COND (IN xa s) (iterate x s f) (x (f xa) (iterate x s f)))" |
|
7714 |
by (import hollight ITERATE_CLAUSES) |
|
7715 |
||
| 19093 | 7716 |
lemma ITERATE_UNION: "ALL u_4247::'q_49843::type => 'q_49843::type => 'q_49843::type. |
7717 |
monoidal u_4247 --> |
|
7718 |
(ALL (f::'q_49828::type => 'q_49843::type) (s::'q_49828::type => bool) |
|
7719 |
x::'q_49828::type => bool. |
|
| 17644 | 7720 |
FINITE s & FINITE x & DISJOINT s x --> |
| 19093 | 7721 |
iterate u_4247 (hollight.UNION s x) f = |
7722 |
u_4247 (iterate u_4247 s f) (iterate u_4247 x f))" |
|
| 17644 | 7723 |
by (import hollight ITERATE_UNION) |
7724 |
||
| 19093 | 7725 |
lemma ITERATE_UNION_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7726 |
monoidal u_4247 --> |
|
| 17644 | 7727 |
(ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. |
| 19093 | 7728 |
FINITE (support u_4247 f s) & |
7729 |
FINITE (support u_4247 f t) & |
|
7730 |
DISJOINT (support u_4247 f s) (support u_4247 f t) --> |
|
7731 |
iterate u_4247 (hollight.UNION s t) f = |
|
7732 |
u_4247 (iterate u_4247 s f) (iterate u_4247 t f))" |
|
| 17644 | 7733 |
by (import hollight ITERATE_UNION_GEN) |
7734 |
||
| 19093 | 7735 |
lemma ITERATE_DIFF: "ALL u::'q_49986::type => 'q_49986::type => 'q_49986::type. |
| 17644 | 7736 |
monoidal u --> |
| 19093 | 7737 |
(ALL (f::'q_49982::type => 'q_49986::type) (s::'q_49982::type => bool) |
7738 |
t::'q_49982::type => bool. |
|
| 17644 | 7739 |
FINITE s & SUBSET t s --> |
7740 |
u (iterate u (DIFF s t) f) (iterate u t f) = iterate u s f)" |
|
7741 |
by (import hollight ITERATE_DIFF) |
|
7742 |
||
| 19093 | 7743 |
lemma ITERATE_DIFF_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7744 |
monoidal u_4247 --> |
|
| 17644 | 7745 |
(ALL (f::'A::type => 'B::type) (s::'A::type => bool) t::'A::type => bool. |
| 19093 | 7746 |
FINITE (support u_4247 f s) & |
7747 |
SUBSET (support u_4247 f t) (support u_4247 f s) --> |
|
7748 |
u_4247 (iterate u_4247 (DIFF s t) f) (iterate u_4247 t f) = |
|
7749 |
iterate u_4247 s f)" |
|
| 17644 | 7750 |
by (import hollight ITERATE_DIFF_GEN) |
7751 |
||
| 19093 | 7752 |
lemma ITERATE_CLOSED: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7753 |
monoidal u_4247 --> |
|
| 17644 | 7754 |
(ALL P::'B::type => bool. |
| 19093 | 7755 |
P (neutral u_4247) & |
7756 |
(ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) --> |
|
| 17644 | 7757 |
(ALL (f::'A::type => 'B::type) x::'A::type => bool. |
7758 |
FINITE x & (ALL xa::'A::type. IN xa x --> P (f xa)) --> |
|
| 19093 | 7759 |
P (iterate u_4247 x f)))" |
| 17644 | 7760 |
by (import hollight ITERATE_CLOSED) |
7761 |
||
| 19093 | 7762 |
lemma ITERATE_CLOSED_GEN: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7763 |
monoidal u_4247 --> |
|
| 17644 | 7764 |
(ALL P::'B::type => bool. |
| 19093 | 7765 |
P (neutral u_4247) & |
7766 |
(ALL (x::'B::type) y::'B::type. P x & P y --> P (u_4247 x y)) --> |
|
| 17644 | 7767 |
(ALL (f::'A::type => 'B::type) s::'A::type => bool. |
| 19093 | 7768 |
FINITE (support u_4247 f s) & |
7769 |
(ALL x::'A::type. IN x s & f x ~= neutral u_4247 --> P (f x)) --> |
|
7770 |
P (iterate u_4247 s f)))" |
|
| 17644 | 7771 |
by (import hollight ITERATE_CLOSED_GEN) |
7772 |
||
| 19093 | 7773 |
lemma ITERATE_RELATED: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7774 |
monoidal u_4247 --> |
|
| 17644 | 7775 |
(ALL R::'B::type => 'B::type => bool. |
| 19093 | 7776 |
R (neutral u_4247) (neutral u_4247) & |
| 17644 | 7777 |
(ALL (x1::'B::type) (y1::'B::type) (x2::'B::type) y2::'B::type. |
| 19093 | 7778 |
R x1 x2 & R y1 y2 --> R (u_4247 x1 y1) (u_4247 x2 y2)) --> |
| 17644 | 7779 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) |
7780 |
x::'A::type => bool. |
|
7781 |
FINITE x & (ALL xa::'A::type. IN xa x --> R (f xa) (g xa)) --> |
|
| 19093 | 7782 |
R (iterate u_4247 x f) (iterate u_4247 x g)))" |
| 17644 | 7783 |
by (import hollight ITERATE_RELATED) |
7784 |
||
| 19093 | 7785 |
lemma ITERATE_EQ_NEUTRAL: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7786 |
monoidal u_4247 --> |
|
| 17644 | 7787 |
(ALL (f::'A::type => 'B::type) s::'A::type => bool. |
| 19093 | 7788 |
(ALL x::'A::type. IN x s --> f x = neutral u_4247) --> |
7789 |
iterate u_4247 s f = neutral u_4247)" |
|
| 17644 | 7790 |
by (import hollight ITERATE_EQ_NEUTRAL) |
7791 |
||
7792 |
lemma ITERATE_SING: "ALL x::'B::type => 'B::type => 'B::type. |
|
7793 |
monoidal x --> |
|
7794 |
(ALL (f::'A::type => 'B::type) xa::'A::type. |
|
7795 |
iterate x (INSERT xa EMPTY) f = f xa)" |
|
7796 |
by (import hollight ITERATE_SING) |
|
7797 |
||
| 19093 | 7798 |
lemma ITERATE_DELTA: "ALL u_4247::'q_50438::type => 'q_50438::type => 'q_50438::type. |
7799 |
monoidal u_4247 --> |
|
7800 |
(ALL (x::'q_50457::type => 'q_50438::type) (xa::'q_50457::type) |
|
7801 |
xb::'q_50457::type => bool. |
|
7802 |
iterate u_4247 xb |
|
7803 |
(%xb::'q_50457::type. COND (xb = xa) (x xb) (neutral u_4247)) = |
|
7804 |
COND (IN xa xb) (x xa) (neutral u_4247))" |
|
| 17644 | 7805 |
by (import hollight ITERATE_DELTA) |
7806 |
||
| 19093 | 7807 |
lemma ITERATE_IMAGE: "ALL u_4247::'q_50536::type => 'q_50536::type => 'q_50536::type. |
7808 |
monoidal u_4247 --> |
|
7809 |
(ALL (f::'q_50535::type => 'q_50507::type) |
|
7810 |
(g::'q_50507::type => 'q_50536::type) x::'q_50535::type => bool. |
|
| 17644 | 7811 |
FINITE x & |
| 19093 | 7812 |
(ALL (xa::'q_50535::type) y::'q_50535::type. |
| 17644 | 7813 |
IN xa x & IN y x & f xa = f y --> xa = y) --> |
| 19093 | 7814 |
iterate u_4247 (IMAGE f x) g = iterate u_4247 x (g o f))" |
| 17644 | 7815 |
by (import hollight ITERATE_IMAGE) |
7816 |
||
| 19093 | 7817 |
lemma ITERATE_BIJECTION: "ALL u_4247::'B::type => 'B::type => 'B::type. |
7818 |
monoidal u_4247 --> |
|
7819 |
(ALL (f::'A::type => 'B::type) (p::'A::type => 'A::type) |
|
7820 |
s::'A::type => bool. |
|
7821 |
FINITE s & |
|
7822 |
(ALL x::'A::type. IN x s --> IN (p x) s) & |
|
7823 |
(ALL y::'A::type. IN y s --> (EX! x::'A::type. IN x s & p x = y)) --> |
|
7824 |
iterate u_4247 s f = iterate u_4247 s (f o p))" |
|
7825 |
by (import hollight ITERATE_BIJECTION) |
|
7826 |
||
7827 |
lemma ITERATE_ITERATE_PRODUCT: "ALL u_4247::'C::type => 'C::type => 'C::type. |
|
7828 |
monoidal u_4247 --> |
|
7829 |
(ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) |
|
7830 |
xb::'A::type => 'B::type => 'C::type. |
|
7831 |
FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> |
|
7832 |
iterate u_4247 x (%i::'A::type. iterate u_4247 (xa i) (xb i)) = |
|
7833 |
iterate u_4247 |
|
7834 |
(GSPEC |
|
7835 |
(%u::'A::type * 'B::type. |
|
7836 |
EX (i::'A::type) j::'B::type. |
|
7837 |
SETSPEC u (IN i x & IN j (xa i)) (i, j))) |
|
7838 |
(GABS |
|
7839 |
(%f::'A::type * 'B::type => 'C::type. |
|
7840 |
ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j))))" |
|
7841 |
by (import hollight ITERATE_ITERATE_PRODUCT) |
|
7842 |
||
7843 |
lemma ITERATE_EQ: "ALL u_4247::'q_50867::type => 'q_50867::type => 'q_50867::type. |
|
7844 |
monoidal u_4247 --> |
|
7845 |
(ALL (f::'q_50871::type => 'q_50867::type) |
|
7846 |
(g::'q_50871::type => 'q_50867::type) x::'q_50871::type => bool. |
|
7847 |
FINITE x & (ALL xa::'q_50871::type. IN xa x --> f xa = g xa) --> |
|
7848 |
iterate u_4247 x f = iterate u_4247 x g)" |
|
7849 |
by (import hollight ITERATE_EQ) |
|
7850 |
||
7851 |
lemma ITERATE_EQ_GENERAL: "ALL u_4247::'C::type => 'C::type => 'C::type. |
|
7852 |
monoidal u_4247 --> |
|
7853 |
(ALL (s::'A::type => bool) (t::'B::type => bool) |
|
7854 |
(f::'A::type => 'C::type) (g::'B::type => 'C::type) |
|
7855 |
h::'A::type => 'B::type. |
|
7856 |
FINITE s & |
|
7857 |
(ALL y::'B::type. IN y t --> (EX! x::'A::type. IN x s & h x = y)) & |
|
7858 |
(ALL x::'A::type. IN x s --> IN (h x) t & g (h x) = f x) --> |
|
7859 |
iterate u_4247 s f = iterate u_4247 t g)" |
|
7860 |
by (import hollight ITERATE_EQ_GENERAL) |
|
7861 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
7862 |
definition nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" where
|
| 19093 | 7863 |
"(op ==::(('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
7864 |
=> (('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
|
| 17644 | 7865 |
=> prop) |
| 19093 | 7866 |
(nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
| 17644 | 7867 |
((iterate::(nat => nat => nat) |
| 19093 | 7868 |
=> ('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7869 |
(plus::nat => nat => nat))" |
| 17644 | 7870 |
|
| 19093 | 7871 |
lemma DEF_nsum: "(op =::(('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
7872 |
=> (('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
|
| 17644 | 7873 |
=> bool) |
| 19093 | 7874 |
(nsum::('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
| 17644 | 7875 |
((iterate::(nat => nat => nat) |
| 19093 | 7876 |
=> ('q_51017::type => bool) => ('q_51017::type => nat) => nat)
|
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7877 |
(plus::nat => nat => nat))" |
| 17644 | 7878 |
by (import hollight DEF_nsum) |
7879 |
||
| 17652 | 7880 |
lemma NEUTRAL_ADD: "(op =::nat => nat => bool) |
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7881 |
((neutral::(nat => nat => nat) => nat) (plus::nat => nat => nat)) (0::nat)" |
| 17644 | 7882 |
by (import hollight NEUTRAL_ADD) |
7883 |
||
| 17652 | 7884 |
lemma NEUTRAL_MUL: "neutral op * = NUMERAL_BIT1 0" |
| 17644 | 7885 |
by (import hollight NEUTRAL_MUL) |
7886 |
||
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7887 |
lemma MONOIDAL_ADD: "(monoidal::(nat => nat => nat) => bool) (plus::nat => nat => nat)" |
| 17644 | 7888 |
by (import hollight MONOIDAL_ADD) |
7889 |
||
7890 |
lemma MONOIDAL_MUL: "(monoidal::(nat => nat => nat) => bool) (op *::nat => nat => nat)" |
|
7891 |
by (import hollight MONOIDAL_MUL) |
|
7892 |
||
| 19093 | 7893 |
lemma NSUM_CLAUSES: "(ALL x::'q_51055::type => nat. nsum EMPTY x = 0) & |
7894 |
(ALL (x::'q_51094::type) (xa::'q_51094::type => nat) |
|
7895 |
xb::'q_51094::type => bool. |
|
| 17644 | 7896 |
FINITE xb --> |
7897 |
nsum (INSERT x xb) xa = COND (IN x xb) (nsum xb xa) (xa x + nsum xb xa))" |
|
7898 |
by (import hollight NSUM_CLAUSES) |
|
7899 |
||
| 19093 | 7900 |
lemma NSUM_UNION: "ALL (x::'q_51120::type => nat) (xa::'q_51120::type => bool) |
7901 |
xb::'q_51120::type => bool. |
|
| 17644 | 7902 |
FINITE xa & FINITE xb & DISJOINT xa xb --> |
7903 |
nsum (hollight.UNION xa xb) x = nsum xa x + nsum xb x" |
|
7904 |
by (import hollight NSUM_UNION) |
|
7905 |
||
| 19093 | 7906 |
lemma NSUM_DIFF: "ALL (f::'q_51175::type => nat) (s::'q_51175::type => bool) |
7907 |
t::'q_51175::type => bool. |
|
| 17644 | 7908 |
FINITE s & SUBSET t s --> nsum (DIFF s t) f = nsum s f - nsum t f" |
7909 |
by (import hollight NSUM_DIFF) |
|
7910 |
||
| 19093 | 7911 |
lemma NSUM_SUPPORT: "ALL (x::'q_51214::type => nat) xa::'q_51214::type => bool. |
|
34974
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
haftmann
parents:
19233
diff
changeset
|
7912 |
FINITE (support plus x xa) --> nsum (support plus x xa) x = nsum xa x" |
| 17644 | 7913 |
by (import hollight NSUM_SUPPORT) |
7914 |
||
| 19093 | 7915 |
lemma NSUM_ADD: "ALL (f::'q_51260::type => nat) (g::'q_51260::type => nat) |
7916 |
s::'q_51260::type => bool. |
|
7917 |
FINITE s --> nsum s (%x::'q_51260::type. f x + g x) = nsum s f + nsum s g" |
|
| 17644 | 7918 |
by (import hollight NSUM_ADD) |
7919 |
||
| 19093 | 7920 |
lemma NSUM_CMUL: "ALL (f::'q_51298::type => nat) (c::nat) s::'q_51298::type => bool. |
7921 |
FINITE s --> nsum s (%x::'q_51298::type. c * f x) = c * nsum s f" |
|
| 17644 | 7922 |
by (import hollight NSUM_CMUL) |
7923 |
||
| 19093 | 7924 |
lemma NSUM_LE: "ALL (x::'q_51337::type => nat) (xa::'q_51337::type => nat) |
7925 |
xb::'q_51337::type => bool. |
|
7926 |
FINITE xb & (ALL xc::'q_51337::type. IN xc xb --> <= (x xc) (xa xc)) --> |
|
| 17644 | 7927 |
<= (nsum xb x) (nsum xb xa)" |
7928 |
by (import hollight NSUM_LE) |
|
7929 |
||
7930 |
lemma NSUM_LT: "ALL (f::'A::type => nat) (g::'A::type => nat) s::'A::type => bool. |
|
7931 |
FINITE s & |
|
7932 |
(ALL x::'A::type. IN x s --> <= (f x) (g x)) & |
|
7933 |
(EX x::'A::type. IN x s & < (f x) (g x)) --> |
|
7934 |
< (nsum s f) (nsum s g)" |
|
7935 |
by (import hollight NSUM_LT) |
|
7936 |
||
| 19093 | 7937 |
lemma NSUM_LT_ALL: "ALL (f::'q_51459::type => nat) (g::'q_51459::type => nat) |
7938 |
s::'q_51459::type => bool. |
|
| 17644 | 7939 |
FINITE s & |
| 19093 | 7940 |
s ~= EMPTY & (ALL x::'q_51459::type. IN x s --> < (f x) (g x)) --> |
| 17644 | 7941 |
< (nsum s f) (nsum s g)" |
7942 |
by (import hollight NSUM_LT_ALL) |
|
7943 |
||
| 19093 | 7944 |
lemma NSUM_EQ: "ALL (x::'q_51501::type => nat) (xa::'q_51501::type => nat) |
7945 |
xb::'q_51501::type => bool. |
|
7946 |
FINITE xb & (ALL xc::'q_51501::type. IN xc xb --> x xc = xa xc) --> |
|
| 17644 | 7947 |
nsum xb x = nsum xb xa" |
7948 |
by (import hollight NSUM_EQ) |
|
7949 |
||
| 19093 | 7950 |
lemma NSUM_CONST: "ALL (c::nat) s::'q_51531::type => bool. |
7951 |
FINITE s --> nsum s (%n::'q_51531::type. c) = CARD s * c" |
|
| 17644 | 7952 |
by (import hollight NSUM_CONST) |
7953 |
||
7954 |
lemma NSUM_EQ_0: "ALL (x::'A::type => nat) xa::'A::type => bool. |
|
| 17652 | 7955 |
(ALL xb::'A::type. IN xb xa --> x xb = 0) --> nsum xa x = 0" |
| 17644 | 7956 |
by (import hollight NSUM_EQ_0) |
7957 |
||
| 17652 | 7958 |
lemma NSUM_0: "ALL x::'A::type => bool. nsum x (%n::'A::type. 0) = 0" |
| 17644 | 7959 |
by (import hollight NSUM_0) |
7960 |
||
| 19093 | 7961 |
lemma NSUM_POS_LE: "ALL (x::'q_51610::type => nat) xa::'q_51610::type => bool. |
7962 |
FINITE xa & (ALL xb::'q_51610::type. IN xb xa --> <= 0 (x xb)) --> |
|
| 17652 | 7963 |
<= 0 (nsum xa x)" |
| 17644 | 7964 |
by (import hollight NSUM_POS_LE) |
7965 |
||
7966 |
lemma NSUM_POS_BOUND: "ALL (f::'A::type => nat) (b::nat) x::'A::type => bool. |
|
7967 |
FINITE x & |
|
| 17652 | 7968 |
(ALL xa::'A::type. IN xa x --> <= 0 (f xa)) & <= (nsum x f) b --> |
| 17644 | 7969 |
(ALL xa::'A::type. IN xa x --> <= (f xa) b)" |
7970 |
by (import hollight NSUM_POS_BOUND) |
|
7971 |
||
| 19093 | 7972 |
lemma NSUM_POS_EQ_0: "ALL (x::'q_51745::type => nat) xa::'q_51745::type => bool. |
| 17644 | 7973 |
FINITE xa & |
| 19093 | 7974 |
(ALL xb::'q_51745::type. IN xb xa --> <= 0 (x xb)) & nsum xa x = 0 --> |
7975 |
(ALL xb::'q_51745::type. IN xb xa --> x xb = 0)" |
|
| 17644 | 7976 |
by (import hollight NSUM_POS_EQ_0) |
7977 |
||
| 19093 | 7978 |
lemma NSUM_SING: "ALL (x::'q_51765::type => nat) xa::'q_51765::type. |
| 17644 | 7979 |
nsum (INSERT xa EMPTY) x = x xa" |
7980 |
by (import hollight NSUM_SING) |
|
7981 |
||
7982 |
lemma NSUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type. |
|
| 17652 | 7983 |
nsum x (%x::'A::type. COND (x = xa) (b::nat) 0) = COND (IN xa x) b 0" |
| 17644 | 7984 |
by (import hollight NSUM_DELTA) |
7985 |
||
7986 |
lemma NSUM_SWAP: "ALL (f::'A::type => 'B::type => nat) (x::'A::type => bool) |
|
7987 |
xa::'B::type => bool. |
|
7988 |
FINITE x & FINITE xa --> |
|
7989 |
nsum x (%i::'A::type. nsum xa (f i)) = |
|
7990 |
nsum xa (%j::'B::type. nsum x (%i::'A::type. f i j))" |
|
7991 |
by (import hollight NSUM_SWAP) |
|
7992 |
||
| 19093 | 7993 |
lemma NSUM_IMAGE: "ALL (x::'q_51905::type => 'q_51881::type) (xa::'q_51881::type => nat) |
7994 |
xb::'q_51905::type => bool. |
|
| 17644 | 7995 |
FINITE xb & |
| 19093 | 7996 |
(ALL (xa::'q_51905::type) y::'q_51905::type. |
| 17644 | 7997 |
IN xa xb & IN y xb & x xa = x y --> xa = y) --> |
7998 |
nsum (IMAGE x xb) xa = nsum xb (xa o x)" |
|
7999 |
by (import hollight NSUM_IMAGE) |
|
8000 |
||
8001 |
lemma NSUM_SUPERSET: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. |
|
8002 |
FINITE u & |
|
| 17652 | 8003 |
SUBSET u v & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) --> |
| 17644 | 8004 |
nsum v f = nsum u f" |
8005 |
by (import hollight NSUM_SUPERSET) |
|
8006 |
||
8007 |
lemma NSUM_UNION_RZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. |
|
| 17652 | 8008 |
FINITE u & (ALL x::'A::type. IN x v & ~ IN x u --> f x = 0) --> |
| 17644 | 8009 |
nsum (hollight.UNION u v) f = nsum u f" |
8010 |
by (import hollight NSUM_UNION_RZERO) |
|
8011 |
||
8012 |
lemma NSUM_UNION_LZERO: "ALL (f::'A::type => nat) (u::'A::type => bool) v::'A::type => bool. |
|
| 17652 | 8013 |
FINITE v & (ALL x::'A::type. IN x u & ~ IN x v --> f x = 0) --> |
| 17644 | 8014 |
nsum (hollight.UNION u v) f = nsum v f" |
8015 |
by (import hollight NSUM_UNION_LZERO) |
|
8016 |
||
| 19093 | 8017 |
lemma NSUM_RESTRICT: "ALL (f::'q_52126::type => nat) s::'q_52126::type => bool. |
| 17644 | 8018 |
FINITE s --> |
| 19093 | 8019 |
nsum s (%x::'q_52126::type. COND (IN x s) (f x) 0) = nsum s f" |
| 17644 | 8020 |
by (import hollight NSUM_RESTRICT) |
8021 |
||
8022 |
lemma NSUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => nat) xb::nat. |
|
8023 |
FINITE x & (ALL xc::'A::type. IN xc x --> <= (xa xc) xb) --> |
|
8024 |
<= (nsum x xa) (CARD x * xb)" |
|
8025 |
by (import hollight NSUM_BOUND) |
|
8026 |
||
| 19093 | 8027 |
lemma NSUM_BOUND_GEN: "ALL (x::'A::type => bool) (xa::'q_52201::type) b::nat. |
| 17644 | 8028 |
FINITE x & |
8029 |
x ~= EMPTY & |
|
8030 |
(ALL xa::'A::type. |
|
8031 |
IN xa x --> <= ((f::'A::type => nat) xa) (DIV b (CARD x))) --> |
|
8032 |
<= (nsum x f) b" |
|
8033 |
by (import hollight NSUM_BOUND_GEN) |
|
8034 |
||
8035 |
lemma NSUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => nat) b::nat. |
|
8036 |
FINITE s & |
|
8037 |
(ALL x::'A::type. IN x s --> <= (f x) b) & |
|
8038 |
(EX x::'A::type. IN x s & < (f x) b) --> |
|
8039 |
< (nsum s f) (CARD s * b)" |
|
8040 |
by (import hollight NSUM_BOUND_LT) |
|
8041 |
||
| 19093 | 8042 |
lemma NSUM_BOUND_LT_ALL: "ALL (s::'q_52344::type => bool) (f::'q_52344::type => nat) b::nat. |
8043 |
FINITE s & s ~= EMPTY & (ALL x::'q_52344::type. IN x s --> < (f x) b) --> |
|
| 17644 | 8044 |
< (nsum s f) (CARD s * b)" |
8045 |
by (import hollight NSUM_BOUND_LT_ALL) |
|
8046 |
||
| 19093 | 8047 |
lemma NSUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_52386::type) b::nat. |
| 17644 | 8048 |
FINITE s & |
8049 |
s ~= EMPTY & |
|
8050 |
(ALL x::'A::type. |
|
8051 |
IN x s --> < ((f::'A::type => nat) x) (DIV b (CARD s))) --> |
|
8052 |
< (nsum s f) b" |
|
8053 |
by (import hollight NSUM_BOUND_LT_GEN) |
|
8054 |
||
| 19093 | 8055 |
lemma NSUM_UNION_EQ: "ALL (s::'q_52445::type => bool) (t::'q_52445::type => bool) |
8056 |
u::'q_52445::type => bool. |
|
| 17644 | 8057 |
FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> |
| 19093 | 8058 |
nsum s (f::'q_52445::type => nat) + nsum t f = nsum u f" |
| 17644 | 8059 |
by (import hollight NSUM_UNION_EQ) |
8060 |
||
8061 |
lemma NSUM_EQ_SUPERSET: "ALL (f::'A::type => nat) (s::'A::type => bool) t::'A::type => bool. |
|
8062 |
FINITE t & |
|
8063 |
SUBSET t s & |
|
8064 |
(ALL x::'A::type. IN x t --> f x = (g::'A::type => nat) x) & |
|
| 17652 | 8065 |
(ALL x::'A::type. IN x s & ~ IN x t --> f x = 0) --> |
| 17644 | 8066 |
nsum s f = nsum t g" |
8067 |
by (import hollight NSUM_EQ_SUPERSET) |
|
8068 |
||
| 19093 | 8069 |
lemma NSUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => nat) r::'q_52556::type. |
| 17644 | 8070 |
FINITE s --> |
8071 |
nsum |
|
8072 |
(GSPEC |
|
8073 |
(%u::'A::type. |
|
8074 |
EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x)) |
|
8075 |
f = |
|
| 17652 | 8076 |
nsum s (%x::'A::type. COND (P x) (f x) 0)" |
| 17644 | 8077 |
by (import hollight NSUM_RESTRICT_SET) |
8078 |
||
| 19093 | 8079 |
lemma NSUM_NSUM_RESTRICT: "ALL (R::'q_52685::type => 'q_52684::type => bool) |
8080 |
(f::'q_52685::type => 'q_52684::type => nat) (s::'q_52685::type => bool) |
|
8081 |
t::'q_52684::type => bool. |
|
| 17644 | 8082 |
FINITE s & FINITE t --> |
8083 |
nsum s |
|
| 19093 | 8084 |
(%x::'q_52685::type. |
| 17644 | 8085 |
nsum |
8086 |
(GSPEC |
|
| 19093 | 8087 |
(%u::'q_52684::type. |
8088 |
EX y::'q_52684::type. SETSPEC u (IN y t & R x y) y)) |
|
| 17644 | 8089 |
(f x)) = |
8090 |
nsum t |
|
| 19093 | 8091 |
(%y::'q_52684::type. |
| 17644 | 8092 |
nsum |
8093 |
(GSPEC |
|
| 19093 | 8094 |
(%u::'q_52685::type. |
8095 |
EX x::'q_52685::type. SETSPEC u (IN x s & R x y) x)) |
|
8096 |
(%x::'q_52685::type. f x y))" |
|
| 17644 | 8097 |
by (import hollight NSUM_NSUM_RESTRICT) |
8098 |
||
| 19093 | 8099 |
lemma CARD_EQ_NSUM: "ALL x::'q_52704::type => bool. |
8100 |
FINITE x --> CARD x = nsum x (%x::'q_52704::type. NUMERAL_BIT1 0)" |
|
| 17644 | 8101 |
by (import hollight CARD_EQ_NSUM) |
8102 |
||
8103 |
lemma NSUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) |
|
8104 |
(t::'B::type => bool) k::'B::type => nat. |
|
8105 |
FINITE s & |
|
8106 |
FINITE t & |
|
8107 |
(ALL j::'B::type. |
|
8108 |
IN j t --> |
|
8109 |
CARD |
|
8110 |
(GSPEC |
|
8111 |
(%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = |
|
8112 |
k j) --> |
|
8113 |
nsum s |
|
8114 |
(%i::'A::type. |
|
8115 |
CARD |
|
8116 |
(GSPEC |
|
8117 |
(%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) = |
|
8118 |
nsum t k" |
|
8119 |
by (import hollight NSUM_MULTICOUNT_GEN) |
|
8120 |
||
8121 |
lemma NSUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) |
|
8122 |
(t::'B::type => bool) k::nat. |
|
8123 |
FINITE s & |
|
8124 |
FINITE t & |
|
8125 |
(ALL j::'B::type. |
|
8126 |
IN j t --> |
|
8127 |
CARD |
|
8128 |
(GSPEC |
|
8129 |
(%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = |
|
8130 |
k) --> |
|
8131 |
nsum s |
|
8132 |
(%i::'A::type. |
|
8133 |
CARD |
|
8134 |
(GSPEC |
|
8135 |
(%u::'B::type. EX j::'B::type. SETSPEC u (IN j t & R i j) j))) = |
|
8136 |
k * CARD t" |
|
8137 |
by (import hollight NSUM_MULTICOUNT) |
|
8138 |
||
8139 |
lemma NSUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => nat) s::'A::type => bool. |
|
8140 |
FINITE s --> |
|
8141 |
nsum s g = |
|
8142 |
nsum (IMAGE f s) |
|
8143 |
(%y::'B::type. |
|
8144 |
nsum |
|
8145 |
(GSPEC |
|
8146 |
(%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x)) |
|
8147 |
g)" |
|
8148 |
by (import hollight NSUM_IMAGE_GEN) |
|
8149 |
||
8150 |
lemma NSUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool) f::'A::type => nat. |
|
| 17652 | 8151 |
FINITE u & FINITE v & (ALL x::'A::type. IN x (DIFF u v) --> f x = 0) --> |
| 17644 | 8152 |
<= (nsum u f) (nsum v f)" |
8153 |
by (import hollight NSUM_SUBSET) |
|
8154 |
||
| 19093 | 8155 |
lemma NSUM_SUBSET_SIMPLE: "ALL (u::'q_53164::type => bool) (v::'q_53164::type => bool) |
8156 |
f::'q_53164::type => nat. |
|
| 17644 | 8157 |
FINITE v & SUBSET u v --> <= (nsum u f) (nsum v f)" |
8158 |
by (import hollight NSUM_SUBSET_SIMPLE) |
|
8159 |
||
8160 |
lemma NSUM_ADD_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat. |
|
8161 |
nsum (dotdot xb xc) (%i::nat. x i + xa i) = |
|
8162 |
nsum (dotdot xb xc) x + nsum (dotdot xb xc) xa" |
|
8163 |
by (import hollight NSUM_ADD_NUMSEG) |
|
8164 |
||
8165 |
lemma NSUM_CMUL_NUMSEG: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat. |
|
8166 |
nsum (dotdot xb xc) (%i::nat. xa * x i) = xa * nsum (dotdot xb xc) x" |
|
8167 |
by (import hollight NSUM_CMUL_NUMSEG) |
|
8168 |
||
8169 |
lemma NSUM_LE_NUMSEG: "ALL (x::nat => nat) (xa::nat => nat) (xb::nat) xc::nat. |
|
8170 |
(ALL i::nat. <= xb i & <= i xc --> <= (x i) (xa i)) --> |
|
8171 |
<= (nsum (dotdot xb xc) x) (nsum (dotdot xb xc) xa)" |
|
8172 |
by (import hollight NSUM_LE_NUMSEG) |
|
8173 |
||
8174 |
lemma NSUM_EQ_NUMSEG: "ALL (f::nat => nat) (g::nat => nat) (m::nat) n::nat. |
|
8175 |
(ALL i::nat. <= m i & <= i n --> f i = g i) --> |
|
8176 |
nsum (dotdot m n) f = nsum (dotdot m n) g" |
|
8177 |
by (import hollight NSUM_EQ_NUMSEG) |
|
8178 |
||
8179 |
lemma NSUM_CONST_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat. |
|
| 17652 | 8180 |
nsum (dotdot xa xb) (%n::nat. x) = (xb + NUMERAL_BIT1 0 - xa) * x" |
| 17644 | 8181 |
by (import hollight NSUM_CONST_NUMSEG) |
8182 |
||
| 19093 | 8183 |
lemma NSUM_EQ_0_NUMSEG: "ALL (x::nat => nat) xa::'q_53403::type. |
| 17652 | 8184 |
(ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = 0) --> |
8185 |
nsum (dotdot m n) x = 0" |
|
| 17644 | 8186 |
by (import hollight NSUM_EQ_0_NUMSEG) |
8187 |
||
| 17652 | 8188 |
lemma NSUM_TRIV_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat. < n m --> nsum (dotdot m n) f = 0" |
| 17644 | 8189 |
by (import hollight NSUM_TRIV_NUMSEG) |
8190 |
||
8191 |
lemma NSUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => nat. |
|
| 17652 | 8192 |
(ALL p::nat. <= x p & <= p xa --> <= 0 (xb p)) --> |
8193 |
<= 0 (nsum (dotdot x xa) xb)" |
|
| 17644 | 8194 |
by (import hollight NSUM_POS_LE_NUMSEG) |
8195 |
||
8196 |
lemma NSUM_POS_EQ_0_NUMSEG: "ALL (f::nat => nat) (m::nat) n::nat. |
|
| 17652 | 8197 |
(ALL p::nat. <= m p & <= p n --> <= 0 (f p)) & |
8198 |
nsum (dotdot m n) f = 0 --> |
|
8199 |
(ALL p::nat. <= m p & <= p n --> f p = 0)" |
|
| 17644 | 8200 |
by (import hollight NSUM_POS_EQ_0_NUMSEG) |
8201 |
||
8202 |
lemma NSUM_SING_NUMSEG: "ALL (x::nat => nat) xa::nat. nsum (dotdot xa xa) x = x xa" |
|
8203 |
by (import hollight NSUM_SING_NUMSEG) |
|
8204 |
||
| 17652 | 8205 |
lemma NSUM_CLAUSES_NUMSEG: "(ALL x::nat. nsum (dotdot x 0) (f::nat => nat) = COND (x = 0) (f 0) 0) & |
| 17644 | 8206 |
(ALL (x::nat) xa::nat. |
8207 |
nsum (dotdot x (Suc xa)) f = |
|
8208 |
COND (<= x (Suc xa)) (nsum (dotdot x xa) f + f (Suc xa)) |
|
8209 |
(nsum (dotdot x xa) f))" |
|
8210 |
by (import hollight NSUM_CLAUSES_NUMSEG) |
|
8211 |
||
8212 |
lemma NSUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => nat. |
|
8213 |
nsum (dotdot a b) (%i::nat. nsum (dotdot c d) (f i)) = |
|
8214 |
nsum (dotdot c d) (%j::nat. nsum (dotdot a b) (%i::nat. f i j))" |
|
8215 |
by (import hollight NSUM_SWAP_NUMSEG) |
|
8216 |
||
8217 |
lemma NSUM_ADD_SPLIT: "ALL (x::nat => nat) (xa::nat) (xb::nat) xc::nat. |
|
| 17652 | 8218 |
<= xa (xb + NUMERAL_BIT1 0) --> |
| 17644 | 8219 |
nsum (dotdot xa (xb + xc)) x = |
| 17652 | 8220 |
nsum (dotdot xa xb) x + nsum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x" |
| 17644 | 8221 |
by (import hollight NSUM_ADD_SPLIT) |
8222 |
||
8223 |
lemma NSUM_OFFSET: "ALL (x::nat => nat) (xa::nat) xb::nat. |
|
8224 |
nsum (dotdot (xa + xb) ((n::nat) + xb)) x = |
|
8225 |
nsum (dotdot xa n) (%i::nat. x (i + xb))" |
|
8226 |
by (import hollight NSUM_OFFSET) |
|
8227 |
||
8228 |
lemma NSUM_OFFSET_0: "ALL (x::nat => nat) (xa::nat) xb::nat. |
|
8229 |
<= xa xb --> |
|
| 17652 | 8230 |
nsum (dotdot xa xb) x = nsum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))" |
| 17644 | 8231 |
by (import hollight NSUM_OFFSET_0) |
8232 |
||
8233 |
lemma NSUM_CLAUSES_LEFT: "ALL (x::nat => nat) (xa::nat) xb::nat. |
|
8234 |
<= xa xb --> |
|
| 17652 | 8235 |
nsum (dotdot xa xb) x = x xa + nsum (dotdot (xa + NUMERAL_BIT1 0) xb) x" |
| 17644 | 8236 |
by (import hollight NSUM_CLAUSES_LEFT) |
8237 |
||
8238 |
lemma NSUM_CLAUSES_RIGHT: "ALL (f::nat => nat) (m::nat) n::nat. |
|
| 17652 | 8239 |
< 0 n & <= m n --> |
8240 |
nsum (dotdot m n) f = nsum (dotdot m (n - NUMERAL_BIT1 0)) f + f n" |
|
| 17644 | 8241 |
by (import hollight NSUM_CLAUSES_RIGHT) |
8242 |
||
| 19093 | 8243 |
lemma NSUM_BIJECTION: "ALL (x::'A::type => nat) (xa::'A::type => 'A::type) xb::'A::type => bool. |
8244 |
FINITE xb & |
|
8245 |
(ALL x::'A::type. IN x xb --> IN (xa x) xb) & |
|
8246 |
(ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) --> |
|
8247 |
nsum xb x = nsum xb (x o xa)" |
|
8248 |
by (import hollight NSUM_BIJECTION) |
|
8249 |
||
8250 |
lemma NSUM_NSUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) |
|
8251 |
xb::'A::type => 'B::type => nat. |
|
8252 |
FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> |
|
8253 |
nsum x (%x::'A::type. nsum (xa x) (xb x)) = |
|
8254 |
nsum |
|
8255 |
(GSPEC |
|
8256 |
(%u::'A::type * 'B::type. |
|
8257 |
EX (i::'A::type) j::'B::type. |
|
8258 |
SETSPEC u (IN i x & IN j (xa i)) (i, j))) |
|
8259 |
(GABS |
|
8260 |
(%f::'A::type * 'B::type => nat. |
|
8261 |
ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))" |
|
8262 |
by (import hollight NSUM_NSUM_PRODUCT) |
|
8263 |
||
8264 |
lemma NSUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool) (xb::'A::type => nat) |
|
8265 |
(xc::'B::type => nat) xd::'A::type => 'B::type. |
|
8266 |
FINITE x & |
|
8267 |
(ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) & |
|
8268 |
(ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) --> |
|
8269 |
nsum x xb = nsum xa xc" |
|
8270 |
by (import hollight NSUM_EQ_GENERAL) |
|
8271 |
||
| 17644 | 8272 |
consts |
| 19093 | 8273 |
sum :: "('q_54215 => bool) => ('q_54215 => hollight.real) => hollight.real"
|
| 17644 | 8274 |
|
8275 |
defs |
|
| 19093 | 8276 |
sum_def: "(op ==::(('q_54215::type => bool)
|
8277 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
8278 |
=> (('q_54215::type => bool)
|
|
8279 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8280 |
=> prop) |
| 19093 | 8281 |
(hollight.sum::('q_54215::type => bool)
|
8282 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8283 |
((iterate::(hollight.real => hollight.real => hollight.real) |
| 19093 | 8284 |
=> ('q_54215::type => bool)
|
8285 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8286 |
(real_add::hollight.real => hollight.real => hollight.real))" |
8287 |
||
| 19093 | 8288 |
lemma DEF_sum: "(op =::(('q_54215::type => bool)
|
8289 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
8290 |
=> (('q_54215::type => bool)
|
|
8291 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8292 |
=> bool) |
| 19093 | 8293 |
(hollight.sum::('q_54215::type => bool)
|
8294 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8295 |
((iterate::(hollight.real => hollight.real => hollight.real) |
| 19093 | 8296 |
=> ('q_54215::type => bool)
|
8297 |
=> ('q_54215::type => hollight.real) => hollight.real)
|
|
| 17644 | 8298 |
(real_add::hollight.real => hollight.real => hollight.real))" |
8299 |
by (import hollight DEF_sum) |
|
8300 |
||
| 17652 | 8301 |
lemma NEUTRAL_REAL_ADD: "neutral real_add = real_of_num 0" |
| 17644 | 8302 |
by (import hollight NEUTRAL_REAL_ADD) |
8303 |
||
| 17652 | 8304 |
lemma NEUTRAL_REAL_MUL: "neutral real_mul = real_of_num (NUMERAL_BIT1 0)" |
| 17644 | 8305 |
by (import hollight NEUTRAL_REAL_MUL) |
8306 |
||
8307 |
lemma MONOIDAL_REAL_ADD: "monoidal real_add" |
|
8308 |
by (import hollight MONOIDAL_REAL_ADD) |
|
8309 |
||
8310 |
lemma MONOIDAL_REAL_MUL: "monoidal real_mul" |
|
8311 |
by (import hollight MONOIDAL_REAL_MUL) |
|
8312 |
||
| 19093 | 8313 |
lemma SUM_CLAUSES: "(ALL x::'q_54257::type => hollight.real. |
| 17652 | 8314 |
hollight.sum EMPTY x = real_of_num 0) & |
| 19093 | 8315 |
(ALL (x::'q_54298::type) (xa::'q_54298::type => hollight.real) |
8316 |
xb::'q_54298::type => bool. |
|
| 17644 | 8317 |
FINITE xb --> |
8318 |
hollight.sum (INSERT x xb) xa = |
|
8319 |
COND (IN x xb) (hollight.sum xb xa) |
|
8320 |
(real_add (xa x) (hollight.sum xb xa)))" |
|
8321 |
by (import hollight SUM_CLAUSES) |
|
8322 |
||
| 19093 | 8323 |
lemma SUM_UNION: "ALL (x::'q_54324::type => hollight.real) (xa::'q_54324::type => bool) |
8324 |
xb::'q_54324::type => bool. |
|
| 17644 | 8325 |
FINITE xa & FINITE xb & DISJOINT xa xb --> |
8326 |
hollight.sum (hollight.UNION xa xb) x = |
|
8327 |
real_add (hollight.sum xa x) (hollight.sum xb x)" |
|
8328 |
by (import hollight SUM_UNION) |
|
8329 |
||
| 19093 | 8330 |
lemma SUM_DIFF: "ALL (x::'q_54364::type => hollight.real) (xa::'q_54364::type => bool) |
8331 |
xb::'q_54364::type => bool. |
|
8332 |
FINITE xa & SUBSET xb xa --> |
|
8333 |
hollight.sum (DIFF xa xb) x = |
|
8334 |
real_sub (hollight.sum xa x) (hollight.sum xb x)" |
|
8335 |
by (import hollight SUM_DIFF) |
|
8336 |
||
8337 |
lemma SUM_SUPPORT: "ALL (x::'q_54403::type => hollight.real) xa::'q_54403::type => bool. |
|
| 17644 | 8338 |
FINITE (support real_add x xa) --> |
8339 |
hollight.sum (support real_add x xa) x = hollight.sum xa x" |
|
8340 |
by (import hollight SUM_SUPPORT) |
|
8341 |
||
| 19093 | 8342 |
lemma SUM_ADD: "ALL (f::'q_54449::type => hollight.real) |
8343 |
(g::'q_54449::type => hollight.real) s::'q_54449::type => bool. |
|
8344 |
FINITE s --> |
|
8345 |
hollight.sum s (%x::'q_54449::type. real_add (f x) (g x)) = |
|
8346 |
real_add (hollight.sum s f) (hollight.sum s g)" |
|
8347 |
by (import hollight SUM_ADD) |
|
8348 |
||
8349 |
lemma SUM_CMUL: "ALL (f::'q_54487::type => hollight.real) (c::hollight.real) |
|
8350 |
s::'q_54487::type => bool. |
|
8351 |
FINITE s --> |
|
8352 |
hollight.sum s (%x::'q_54487::type. real_mul c (f x)) = |
|
8353 |
real_mul c (hollight.sum s f)" |
|
8354 |
by (import hollight SUM_CMUL) |
|
8355 |
||
8356 |
lemma SUM_NEG: "ALL (x::'q_54530::type => hollight.real) xa::'q_54530::type => bool. |
|
8357 |
FINITE xa --> |
|
8358 |
hollight.sum xa (%xa::'q_54530::type. real_neg (x xa)) = |
|
8359 |
real_neg (hollight.sum xa x)" |
|
8360 |
by (import hollight SUM_NEG) |
|
8361 |
||
8362 |
lemma SUM_SUB: "ALL (x::'q_54565::type => hollight.real) |
|
8363 |
(xa::'q_54565::type => hollight.real) xb::'q_54565::type => bool. |
|
8364 |
FINITE xb --> |
|
8365 |
hollight.sum xb (%xb::'q_54565::type. real_sub (x xb) (xa xb)) = |
|
8366 |
real_sub (hollight.sum xb x) (hollight.sum xb xa)" |
|
8367 |
by (import hollight SUM_SUB) |
|
8368 |
||
8369 |
lemma SUM_LE: "ALL (x::'q_54607::type => hollight.real) |
|
8370 |
(xa::'q_54607::type => hollight.real) xb::'q_54607::type => bool. |
|
8371 |
FINITE xb & |
|
8372 |
(ALL xc::'q_54607::type. IN xc xb --> real_le (x xc) (xa xc)) --> |
|
8373 |
real_le (hollight.sum xb x) (hollight.sum xb xa)" |
|
8374 |
by (import hollight SUM_LE) |
|
8375 |
||
| 17644 | 8376 |
lemma SUM_LT: "ALL (f::'A::type => hollight.real) (g::'A::type => hollight.real) |
8377 |
s::'A::type => bool. |
|
8378 |
FINITE s & |
|
8379 |
(ALL x::'A::type. IN x s --> real_le (f x) (g x)) & |
|
8380 |
(EX x::'A::type. IN x s & real_lt (f x) (g x)) --> |
|
8381 |
real_lt (hollight.sum s f) (hollight.sum s g)" |
|
8382 |
by (import hollight SUM_LT) |
|
8383 |
||
| 19093 | 8384 |
lemma SUM_LT_ALL: "ALL (f::'q_54729::type => hollight.real) |
8385 |
(g::'q_54729::type => hollight.real) s::'q_54729::type => bool. |
|
| 17644 | 8386 |
FINITE s & |
| 19093 | 8387 |
s ~= EMPTY & (ALL x::'q_54729::type. IN x s --> real_lt (f x) (g x)) --> |
| 17644 | 8388 |
real_lt (hollight.sum s f) (hollight.sum s g)" |
8389 |
by (import hollight SUM_LT_ALL) |
|
8390 |
||
| 19093 | 8391 |
lemma SUM_EQ: "ALL (x::'q_54771::type => hollight.real) |
8392 |
(xa::'q_54771::type => hollight.real) xb::'q_54771::type => bool. |
|
8393 |
FINITE xb & (ALL xc::'q_54771::type. IN xc xb --> x xc = xa xc) --> |
|
8394 |
hollight.sum xb x = hollight.sum xb xa" |
|
8395 |
by (import hollight SUM_EQ) |
|
8396 |
||
8397 |
lemma SUM_ABS: "ALL (f::'q_54830::type => hollight.real) s::'q_54830::type => bool. |
|
8398 |
FINITE s --> |
|
8399 |
real_le (real_abs (hollight.sum s f)) |
|
8400 |
(hollight.sum s (%x::'q_54830::type. real_abs (f x)))" |
|
8401 |
by (import hollight SUM_ABS) |
|
8402 |
||
8403 |
lemma SUM_CONST: "ALL (c::hollight.real) s::'q_54851::type => bool. |
|
8404 |
FINITE s --> |
|
8405 |
hollight.sum s (%n::'q_54851::type. c) = |
|
8406 |
real_mul (real_of_num (CARD s)) c" |
|
8407 |
by (import hollight SUM_CONST) |
|
8408 |
||
8409 |
lemma SUM_EQ_0: "ALL (x::'A::type => hollight.real) xa::'A::type => bool. |
|
8410 |
(ALL xb::'A::type. IN xb xa --> x xb = real_of_num 0) --> |
|
8411 |
hollight.sum xa x = real_of_num 0" |
|
8412 |
by (import hollight SUM_EQ_0) |
|
8413 |
||
8414 |
lemma SUM_0: "ALL x::'A::type => bool. |
|
8415 |
hollight.sum x (%n::'A::type. real_of_num 0) = real_of_num 0" |
|
8416 |
by (import hollight SUM_0) |
|
8417 |
||
8418 |
lemma SUM_POS_LE: "ALL (x::'q_54944::type => hollight.real) xa::'q_54944::type => bool. |
|
| 17644 | 8419 |
FINITE xa & |
| 19093 | 8420 |
(ALL xb::'q_54944::type. IN xb xa --> real_le (real_of_num 0) (x xb)) --> |
| 17652 | 8421 |
real_le (real_of_num 0) (hollight.sum xa x)" |
| 17644 | 8422 |
by (import hollight SUM_POS_LE) |
8423 |
||
8424 |
lemma SUM_POS_BOUND: "ALL (f::'A::type => hollight.real) (b::hollight.real) x::'A::type => bool. |
|
8425 |
FINITE x & |
|
| 17652 | 8426 |
(ALL xa::'A::type. IN xa x --> real_le (real_of_num 0) (f xa)) & |
| 17644 | 8427 |
real_le (hollight.sum x f) b --> |
8428 |
(ALL xa::'A::type. IN xa x --> real_le (f xa) b)" |
|
8429 |
by (import hollight SUM_POS_BOUND) |
|
8430 |
||
| 19093 | 8431 |
lemma SUM_POS_EQ_0: "ALL (x::'q_55091::type => hollight.real) xa::'q_55091::type => bool. |
| 17644 | 8432 |
FINITE xa & |
| 19093 | 8433 |
(ALL xb::'q_55091::type. IN xb xa --> real_le (real_of_num 0) (x xb)) & |
| 17652 | 8434 |
hollight.sum xa x = real_of_num 0 --> |
| 19093 | 8435 |
(ALL xb::'q_55091::type. IN xb xa --> x xb = real_of_num 0)" |
| 17644 | 8436 |
by (import hollight SUM_POS_EQ_0) |
8437 |
||
| 19093 | 8438 |
lemma SUM_SING: "ALL (x::'q_55113::type => hollight.real) xa::'q_55113::type. |
| 17644 | 8439 |
hollight.sum (INSERT xa EMPTY) x = x xa" |
8440 |
by (import hollight SUM_SING) |
|
8441 |
||
8442 |
lemma SUM_DELTA: "ALL (x::'A::type => bool) xa::'A::type. |
|
8443 |
hollight.sum x |
|
| 17652 | 8444 |
(%x::'A::type. COND (x = xa) (b::hollight.real) (real_of_num 0)) = |
8445 |
COND (IN xa x) b (real_of_num 0)" |
|
| 17644 | 8446 |
by (import hollight SUM_DELTA) |
8447 |
||
| 19093 | 8448 |
lemma SUM_SWAP: "ALL (f::'A::type => 'B::type => hollight.real) (x::'A::type => bool) |
8449 |
xa::'B::type => bool. |
|
8450 |
FINITE x & FINITE xa --> |
|
8451 |
hollight.sum x (%i::'A::type. hollight.sum xa (f i)) = |
|
8452 |
hollight.sum xa (%j::'B::type. hollight.sum x (%i::'A::type. f i j))" |
|
8453 |
by (import hollight SUM_SWAP) |
|
8454 |
||
8455 |
lemma SUM_IMAGE: "ALL (x::'q_55257::type => 'q_55233::type) |
|
8456 |
(xa::'q_55233::type => hollight.real) xb::'q_55257::type => bool. |
|
| 17644 | 8457 |
FINITE xb & |
| 19093 | 8458 |
(ALL (xa::'q_55257::type) y::'q_55257::type. |
| 17644 | 8459 |
IN xa xb & IN y xb & x xa = x y --> xa = y) --> |
8460 |
hollight.sum (IMAGE x xb) xa = hollight.sum xb (xa o x)" |
|
8461 |
by (import hollight SUM_IMAGE) |
|
8462 |
||
8463 |
lemma SUM_SUPERSET: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) |
|
8464 |
v::'A::type => bool. |
|
8465 |
FINITE u & |
|
8466 |
SUBSET u v & |
|
| 17652 | 8467 |
(ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) --> |
| 17644 | 8468 |
hollight.sum v f = hollight.sum u f" |
8469 |
by (import hollight SUM_SUPERSET) |
|
8470 |
||
8471 |
lemma SUM_UNION_RZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) |
|
8472 |
v::'A::type => bool. |
|
8473 |
FINITE u & |
|
| 17652 | 8474 |
(ALL x::'A::type. IN x v & ~ IN x u --> f x = real_of_num 0) --> |
| 17644 | 8475 |
hollight.sum (hollight.UNION u v) f = hollight.sum u f" |
8476 |
by (import hollight SUM_UNION_RZERO) |
|
8477 |
||
8478 |
lemma SUM_UNION_LZERO: "ALL (f::'A::type => hollight.real) (u::'A::type => bool) |
|
8479 |
v::'A::type => bool. |
|
8480 |
FINITE v & |
|
| 17652 | 8481 |
(ALL x::'A::type. IN x u & ~ IN x v --> f x = real_of_num 0) --> |
| 17644 | 8482 |
hollight.sum (hollight.UNION u v) f = hollight.sum v f" |
8483 |
by (import hollight SUM_UNION_LZERO) |
|
8484 |
||
| 19093 | 8485 |
lemma SUM_RESTRICT: "ALL (f::'q_55484::type => hollight.real) s::'q_55484::type => bool. |
| 17644 | 8486 |
FINITE s --> |
8487 |
hollight.sum s |
|
| 19093 | 8488 |
(%x::'q_55484::type. COND (IN x s) (f x) (real_of_num 0)) = |
| 17644 | 8489 |
hollight.sum s f" |
8490 |
by (import hollight SUM_RESTRICT) |
|
8491 |
||
| 19093 | 8492 |
lemma SUM_BOUND: "ALL (x::'A::type => bool) (xa::'A::type => hollight.real) xb::hollight.real. |
8493 |
FINITE x & (ALL xc::'A::type. IN xc x --> real_le (xa xc) xb) --> |
|
8494 |
real_le (hollight.sum x xa) (real_mul (real_of_num (CARD x)) xb)" |
|
8495 |
by (import hollight SUM_BOUND) |
|
8496 |
||
8497 |
lemma SUM_BOUND_GEN: "ALL (s::'A::type => bool) (t::'q_55543::type) b::hollight.real. |
|
| 17644 | 8498 |
FINITE s & |
8499 |
s ~= EMPTY & |
|
8500 |
(ALL x::'A::type. |
|
8501 |
IN x s --> |
|
8502 |
real_le ((f::'A::type => hollight.real) x) |
|
8503 |
(real_div b (real_of_num (CARD s)))) --> |
|
8504 |
real_le (hollight.sum s f) b" |
|
8505 |
by (import hollight SUM_BOUND_GEN) |
|
8506 |
||
8507 |
lemma SUM_ABS_BOUND: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real. |
|
8508 |
FINITE s & (ALL x::'A::type. IN x s --> real_le (real_abs (f x)) b) --> |
|
8509 |
real_le (real_abs (hollight.sum s f)) (real_mul (real_of_num (CARD s)) b)" |
|
8510 |
by (import hollight SUM_ABS_BOUND) |
|
8511 |
||
8512 |
lemma SUM_BOUND_LT: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) b::hollight.real. |
|
8513 |
FINITE s & |
|
8514 |
(ALL x::'A::type. IN x s --> real_le (f x) b) & |
|
8515 |
(EX x::'A::type. IN x s & real_lt (f x) b) --> |
|
8516 |
real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" |
|
8517 |
by (import hollight SUM_BOUND_LT) |
|
8518 |
||
| 19093 | 8519 |
lemma SUM_BOUND_LT_ALL: "ALL (s::'q_55748::type => bool) (f::'q_55748::type => hollight.real) |
| 17644 | 8520 |
b::hollight.real. |
8521 |
FINITE s & |
|
| 19093 | 8522 |
s ~= EMPTY & (ALL x::'q_55748::type. IN x s --> real_lt (f x) b) --> |
| 17644 | 8523 |
real_lt (hollight.sum s f) (real_mul (real_of_num (CARD s)) b)" |
8524 |
by (import hollight SUM_BOUND_LT_ALL) |
|
8525 |
||
| 19093 | 8526 |
lemma SUM_BOUND_LT_GEN: "ALL (s::'A::type => bool) (t::'q_55770::type) b::hollight.real. |
| 17644 | 8527 |
FINITE s & |
8528 |
s ~= EMPTY & |
|
8529 |
(ALL x::'A::type. |
|
8530 |
IN x s --> |
|
8531 |
real_lt ((f::'A::type => hollight.real) x) |
|
8532 |
(real_div b (real_of_num (CARD s)))) --> |
|
8533 |
real_lt (hollight.sum s f) b" |
|
8534 |
by (import hollight SUM_BOUND_LT_GEN) |
|
8535 |
||
| 19093 | 8536 |
lemma SUM_UNION_EQ: "ALL (s::'q_55831::type => bool) (t::'q_55831::type => bool) |
8537 |
u::'q_55831::type => bool. |
|
| 17644 | 8538 |
FINITE u & hollight.INTER s t = EMPTY & hollight.UNION s t = u --> |
| 19093 | 8539 |
real_add (hollight.sum s (f::'q_55831::type => hollight.real)) |
| 17644 | 8540 |
(hollight.sum t f) = |
8541 |
hollight.sum u f" |
|
8542 |
by (import hollight SUM_UNION_EQ) |
|
8543 |
||
8544 |
lemma SUM_EQ_SUPERSET: "ALL (f::'A::type => hollight.real) (s::'A::type => bool) |
|
8545 |
t::'A::type => bool. |
|
8546 |
FINITE t & |
|
8547 |
SUBSET t s & |
|
8548 |
(ALL x::'A::type. IN x t --> f x = (g::'A::type => hollight.real) x) & |
|
| 17652 | 8549 |
(ALL x::'A::type. IN x s & ~ IN x t --> f x = real_of_num 0) --> |
| 17644 | 8550 |
hollight.sum s f = hollight.sum t g" |
8551 |
by (import hollight SUM_EQ_SUPERSET) |
|
8552 |
||
| 19093 | 8553 |
lemma SUM_RESTRICT_SET: "ALL (s::'A::type => bool) (f::'A::type => hollight.real) r::'q_55944::type. |
| 17644 | 8554 |
FINITE s --> |
8555 |
hollight.sum |
|
8556 |
(GSPEC |
|
8557 |
(%u::'A::type. |
|
8558 |
EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x)) |
|
8559 |
f = |
|
| 17652 | 8560 |
hollight.sum s (%x::'A::type. COND (P x) (f x) (real_of_num 0))" |
| 17644 | 8561 |
by (import hollight SUM_RESTRICT_SET) |
8562 |
||
| 19093 | 8563 |
lemma SUM_SUM_RESTRICT: "ALL (R::'q_56075::type => 'q_56074::type => bool) |
8564 |
(f::'q_56075::type => 'q_56074::type => hollight.real) |
|
8565 |
(s::'q_56075::type => bool) t::'q_56074::type => bool. |
|
| 17644 | 8566 |
FINITE s & FINITE t --> |
8567 |
hollight.sum s |
|
| 19093 | 8568 |
(%x::'q_56075::type. |
| 17644 | 8569 |
hollight.sum |
8570 |
(GSPEC |
|
| 19093 | 8571 |
(%u::'q_56074::type. |
8572 |
EX y::'q_56074::type. SETSPEC u (IN y t & R x y) y)) |
|
| 17644 | 8573 |
(f x)) = |
8574 |
hollight.sum t |
|
| 19093 | 8575 |
(%y::'q_56074::type. |
| 17644 | 8576 |
hollight.sum |
8577 |
(GSPEC |
|
| 19093 | 8578 |
(%u::'q_56075::type. |
8579 |
EX x::'q_56075::type. SETSPEC u (IN x s & R x y) x)) |
|
8580 |
(%x::'q_56075::type. f x y))" |
|
| 17644 | 8581 |
by (import hollight SUM_SUM_RESTRICT) |
8582 |
||
| 19093 | 8583 |
lemma CARD_EQ_SUM: "ALL x::'q_56096::type => bool. |
| 17644 | 8584 |
FINITE x --> |
8585 |
real_of_num (CARD x) = |
|
| 19093 | 8586 |
hollight.sum x (%x::'q_56096::type. real_of_num (NUMERAL_BIT1 0))" |
| 17644 | 8587 |
by (import hollight CARD_EQ_SUM) |
8588 |
||
8589 |
lemma SUM_MULTICOUNT_GEN: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) |
|
8590 |
(t::'B::type => bool) k::'B::type => nat. |
|
8591 |
FINITE s & |
|
8592 |
FINITE t & |
|
8593 |
(ALL j::'B::type. |
|
8594 |
IN j t --> |
|
8595 |
CARD |
|
8596 |
(GSPEC |
|
8597 |
(%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = |
|
8598 |
k j) --> |
|
8599 |
hollight.sum s |
|
8600 |
(%i::'A::type. |
|
8601 |
real_of_num |
|
8602 |
(CARD |
|
8603 |
(GSPEC |
|
8604 |
(%u::'B::type. |
|
8605 |
EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) = |
|
8606 |
hollight.sum t (%i::'B::type. real_of_num (k i))" |
|
8607 |
by (import hollight SUM_MULTICOUNT_GEN) |
|
8608 |
||
8609 |
lemma SUM_MULTICOUNT: "ALL (R::'A::type => 'B::type => bool) (s::'A::type => bool) |
|
8610 |
(t::'B::type => bool) k::nat. |
|
8611 |
FINITE s & |
|
8612 |
FINITE t & |
|
8613 |
(ALL j::'B::type. |
|
8614 |
IN j t --> |
|
8615 |
CARD |
|
8616 |
(GSPEC |
|
8617 |
(%u::'A::type. EX i::'A::type. SETSPEC u (IN i s & R i j) i)) = |
|
8618 |
k) --> |
|
8619 |
hollight.sum s |
|
8620 |
(%i::'A::type. |
|
8621 |
real_of_num |
|
8622 |
(CARD |
|
8623 |
(GSPEC |
|
8624 |
(%u::'B::type. |
|
8625 |
EX j::'B::type. SETSPEC u (IN j t & R i j) j)))) = |
|
8626 |
real_of_num (k * CARD t)" |
|
8627 |
by (import hollight SUM_MULTICOUNT) |
|
8628 |
||
8629 |
lemma SUM_IMAGE_GEN: "ALL (f::'A::type => 'B::type) (g::'A::type => hollight.real) |
|
8630 |
s::'A::type => bool. |
|
8631 |
FINITE s --> |
|
8632 |
hollight.sum s g = |
|
8633 |
hollight.sum (IMAGE f s) |
|
8634 |
(%y::'B::type. |
|
8635 |
hollight.sum |
|
8636 |
(GSPEC |
|
8637 |
(%u::'A::type. EX x::'A::type. SETSPEC u (IN x s & f x = y) x)) |
|
8638 |
g)" |
|
8639 |
by (import hollight SUM_IMAGE_GEN) |
|
8640 |
||
| 19093 | 8641 |
lemma REAL_OF_NUM_SUM: "ALL (f::'q_56493::type => nat) s::'q_56493::type => bool. |
| 17644 | 8642 |
FINITE s --> |
8643 |
real_of_num (nsum s f) = |
|
| 19093 | 8644 |
hollight.sum s (%x::'q_56493::type. real_of_num (f x))" |
| 17644 | 8645 |
by (import hollight REAL_OF_NUM_SUM) |
8646 |
||
8647 |
lemma SUM_SUBSET: "ALL (u::'A::type => bool) (v::'A::type => bool) |
|
8648 |
f::'A::type => hollight.real. |
|
8649 |
FINITE u & |
|
8650 |
FINITE v & |
|
| 17652 | 8651 |
(ALL x::'A::type. IN x (DIFF u v) --> real_le (f x) (real_of_num 0)) & |
8652 |
(ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) --> |
|
| 17644 | 8653 |
real_le (hollight.sum u f) (hollight.sum v f)" |
8654 |
by (import hollight SUM_SUBSET) |
|
8655 |
||
8656 |
lemma SUM_SUBSET_SIMPLE: "ALL (u::'A::type => bool) (v::'A::type => bool) |
|
8657 |
f::'A::type => hollight.real. |
|
8658 |
FINITE v & |
|
8659 |
SUBSET u v & |
|
| 17652 | 8660 |
(ALL x::'A::type. IN x (DIFF v u) --> real_le (real_of_num 0) (f x)) --> |
| 17644 | 8661 |
real_le (hollight.sum u f) (hollight.sum v f)" |
8662 |
by (import hollight SUM_SUBSET_SIMPLE) |
|
8663 |
||
8664 |
lemma SUM_ADD_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. |
|
8665 |
hollight.sum (dotdot xb xc) (%i::nat. real_add (x i) (xa i)) = |
|
8666 |
real_add (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" |
|
8667 |
by (import hollight SUM_ADD_NUMSEG) |
|
8668 |
||
8669 |
lemma SUM_CMUL_NUMSEG: "ALL (x::nat => hollight.real) (xa::hollight.real) (xb::nat) xc::nat. |
|
8670 |
hollight.sum (dotdot xb xc) (%i::nat. real_mul xa (x i)) = |
|
8671 |
real_mul xa (hollight.sum (dotdot xb xc) x)" |
|
8672 |
by (import hollight SUM_CMUL_NUMSEG) |
|
8673 |
||
8674 |
lemma SUM_NEG_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. |
|
8675 |
hollight.sum (dotdot xa xb) (%i::nat. real_neg (x i)) = |
|
8676 |
real_neg (hollight.sum (dotdot xa xb) x)" |
|
8677 |
by (import hollight SUM_NEG_NUMSEG) |
|
8678 |
||
8679 |
lemma SUM_SUB_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. |
|
8680 |
hollight.sum (dotdot xb xc) (%i::nat. real_sub (x i) (xa i)) = |
|
8681 |
real_sub (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" |
|
8682 |
by (import hollight SUM_SUB_NUMSEG) |
|
8683 |
||
8684 |
lemma SUM_LE_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat => hollight.real) (xb::nat) xc::nat. |
|
8685 |
(ALL i::nat. <= xb i & <= i xc --> real_le (x i) (xa i)) --> |
|
8686 |
real_le (hollight.sum (dotdot xb xc) x) (hollight.sum (dotdot xb xc) xa)" |
|
8687 |
by (import hollight SUM_LE_NUMSEG) |
|
8688 |
||
8689 |
lemma SUM_EQ_NUMSEG: "ALL (f::nat => hollight.real) (g::nat => hollight.real) (m::nat) n::nat. |
|
8690 |
(ALL i::nat. <= m i & <= i n --> f i = g i) --> |
|
8691 |
hollight.sum (dotdot m n) f = hollight.sum (dotdot m n) g" |
|
8692 |
by (import hollight SUM_EQ_NUMSEG) |
|
8693 |
||
8694 |
lemma SUM_ABS_NUMSEG: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. |
|
8695 |
real_le (real_abs (hollight.sum (dotdot xa xb) x)) |
|
8696 |
(hollight.sum (dotdot xa xb) (%i::nat. real_abs (x i)))" |
|
8697 |
by (import hollight SUM_ABS_NUMSEG) |
|
8698 |
||
8699 |
lemma SUM_CONST_NUMSEG: "ALL (x::hollight.real) (xa::nat) xb::nat. |
|
8700 |
hollight.sum (dotdot xa xb) (%n::nat. x) = |
|
| 17652 | 8701 |
real_mul (real_of_num (xb + NUMERAL_BIT1 0 - xa)) x" |
| 17644 | 8702 |
by (import hollight SUM_CONST_NUMSEG) |
8703 |
||
| 19093 | 8704 |
lemma SUM_EQ_0_NUMSEG: "ALL (x::nat => hollight.real) xa::'q_57019::type. |
| 17652 | 8705 |
(ALL i::nat. <= (m::nat) i & <= i (n::nat) --> x i = real_of_num 0) --> |
8706 |
hollight.sum (dotdot m n) x = real_of_num 0" |
|
| 17644 | 8707 |
by (import hollight SUM_EQ_0_NUMSEG) |
8708 |
||
8709 |
lemma SUM_TRIV_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat. |
|
| 17652 | 8710 |
< n m --> hollight.sum (dotdot m n) f = real_of_num 0" |
| 17644 | 8711 |
by (import hollight SUM_TRIV_NUMSEG) |
8712 |
||
8713 |
lemma SUM_POS_LE_NUMSEG: "ALL (x::nat) (xa::nat) xb::nat => hollight.real. |
|
| 17652 | 8714 |
(ALL p::nat. <= x p & <= p xa --> real_le (real_of_num 0) (xb p)) --> |
8715 |
real_le (real_of_num 0) (hollight.sum (dotdot x xa) xb)" |
|
| 17644 | 8716 |
by (import hollight SUM_POS_LE_NUMSEG) |
8717 |
||
8718 |
lemma SUM_POS_EQ_0_NUMSEG: "ALL (f::nat => hollight.real) (m::nat) n::nat. |
|
| 17652 | 8719 |
(ALL p::nat. <= m p & <= p n --> real_le (real_of_num 0) (f p)) & |
8720 |
hollight.sum (dotdot m n) f = real_of_num 0 --> |
|
8721 |
(ALL p::nat. <= m p & <= p n --> f p = real_of_num 0)" |
|
| 17644 | 8722 |
by (import hollight SUM_POS_EQ_0_NUMSEG) |
8723 |
||
8724 |
lemma SUM_SING_NUMSEG: "ALL (x::nat => hollight.real) xa::nat. hollight.sum (dotdot xa xa) x = x xa" |
|
8725 |
by (import hollight SUM_SING_NUMSEG) |
|
8726 |
||
8727 |
lemma SUM_CLAUSES_NUMSEG: "(ALL x::nat. |
|
| 17652 | 8728 |
hollight.sum (dotdot x 0) (f::nat => hollight.real) = |
8729 |
COND (x = 0) (f 0) (real_of_num 0)) & |
|
| 17644 | 8730 |
(ALL (x::nat) xa::nat. |
8731 |
hollight.sum (dotdot x (Suc xa)) f = |
|
8732 |
COND (<= x (Suc xa)) |
|
8733 |
(real_add (hollight.sum (dotdot x xa) f) (f (Suc xa))) |
|
8734 |
(hollight.sum (dotdot x xa) f))" |
|
8735 |
by (import hollight SUM_CLAUSES_NUMSEG) |
|
8736 |
||
8737 |
lemma SUM_SWAP_NUMSEG: "ALL (a::nat) (b::nat) (c::nat) (d::nat) f::nat => nat => hollight.real. |
|
8738 |
hollight.sum (dotdot a b) (%i::nat. hollight.sum (dotdot c d) (f i)) = |
|
8739 |
hollight.sum (dotdot c d) |
|
8740 |
(%j::nat. hollight.sum (dotdot a b) (%i::nat. f i j))" |
|
8741 |
by (import hollight SUM_SWAP_NUMSEG) |
|
8742 |
||
8743 |
lemma SUM_ADD_SPLIT: "ALL (x::nat => hollight.real) (xa::nat) (xb::nat) xc::nat. |
|
| 17652 | 8744 |
<= xa (xb + NUMERAL_BIT1 0) --> |
| 17644 | 8745 |
hollight.sum (dotdot xa (xb + xc)) x = |
8746 |
real_add (hollight.sum (dotdot xa xb) x) |
|
| 17652 | 8747 |
(hollight.sum (dotdot (xb + NUMERAL_BIT1 0) (xb + xc)) x)" |
| 17644 | 8748 |
by (import hollight SUM_ADD_SPLIT) |
8749 |
||
| 19093 | 8750 |
lemma SUM_OFFSET: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. |
8751 |
hollight.sum (dotdot (xa + xb) ((n::nat) + xb)) x = |
|
8752 |
hollight.sum (dotdot xa n) (%i::nat. x (i + xb))" |
|
8753 |
by (import hollight SUM_OFFSET) |
|
8754 |
||
| 17644 | 8755 |
lemma SUM_OFFSET_0: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. |
8756 |
<= xa xb --> |
|
8757 |
hollight.sum (dotdot xa xb) x = |
|
| 17652 | 8758 |
hollight.sum (dotdot 0 (xb - xa)) (%i::nat. x (i + xa))" |
| 17644 | 8759 |
by (import hollight SUM_OFFSET_0) |
8760 |
||
8761 |
lemma SUM_CLAUSES_LEFT: "ALL (x::nat => hollight.real) (xa::nat) xb::nat. |
|
8762 |
<= xa xb --> |
|
8763 |
hollight.sum (dotdot xa xb) x = |
|
| 17652 | 8764 |
real_add (x xa) (hollight.sum (dotdot (xa + NUMERAL_BIT1 0) xb) x)" |
| 17644 | 8765 |
by (import hollight SUM_CLAUSES_LEFT) |
8766 |
||
8767 |
lemma SUM_CLAUSES_RIGHT: "ALL (f::nat => hollight.real) (m::nat) n::nat. |
|
| 17652 | 8768 |
< 0 n & <= m n --> |
| 17644 | 8769 |
hollight.sum (dotdot m n) f = |
| 17652 | 8770 |
real_add (hollight.sum (dotdot m (n - NUMERAL_BIT1 0)) f) (f n)" |
| 17644 | 8771 |
by (import hollight SUM_CLAUSES_RIGHT) |
8772 |
||
8773 |
lemma REAL_OF_NUM_SUM_NUMSEG: "ALL (x::nat => nat) (xa::nat) xb::nat. |
|
8774 |
real_of_num (nsum (dotdot xa xb) x) = |
|
8775 |
hollight.sum (dotdot xa xb) (%i::nat. real_of_num (x i))" |
|
8776 |
by (import hollight REAL_OF_NUM_SUM_NUMSEG) |
|
8777 |
||
| 19093 | 8778 |
lemma SUM_BIJECTION: "ALL (x::'A::type => hollight.real) (xa::'A::type => 'A::type) |
8779 |
xb::'A::type => bool. |
|
8780 |
FINITE xb & |
|
8781 |
(ALL x::'A::type. IN x xb --> IN (xa x) xb) & |
|
8782 |
(ALL y::'A::type. IN y xb --> (EX! x::'A::type. IN x xb & xa x = y)) --> |
|
8783 |
hollight.sum xb x = hollight.sum xb (x o xa)" |
|
8784 |
by (import hollight SUM_BIJECTION) |
|
8785 |
||
8786 |
lemma SUM_SUM_PRODUCT: "ALL (x::'A::type => bool) (xa::'A::type => 'B::type => bool) |
|
8787 |
xb::'A::type => 'B::type => hollight.real. |
|
8788 |
FINITE x & (ALL i::'A::type. IN i x --> FINITE (xa i)) --> |
|
8789 |
hollight.sum x (%x::'A::type. hollight.sum (xa x) (xb x)) = |
|
8790 |
hollight.sum |
|
8791 |
(GSPEC |
|
8792 |
(%u::'A::type * 'B::type. |
|
8793 |
EX (i::'A::type) j::'B::type. |
|
8794 |
SETSPEC u (IN i x & IN j (xa i)) (i, j))) |
|
8795 |
(GABS |
|
8796 |
(%f::'A::type * 'B::type => hollight.real. |
|
8797 |
ALL (i::'A::type) j::'B::type. GEQ (f (i, j)) (xb i j)))" |
|
8798 |
by (import hollight SUM_SUM_PRODUCT) |
|
8799 |
||
8800 |
lemma SUM_EQ_GENERAL: "ALL (x::'A::type => bool) (xa::'B::type => bool) |
|
8801 |
(xb::'A::type => hollight.real) (xc::'B::type => hollight.real) |
|
8802 |
xd::'A::type => 'B::type. |
|
8803 |
FINITE x & |
|
8804 |
(ALL y::'B::type. IN y xa --> (EX! xa::'A::type. IN xa x & xd xa = y)) & |
|
8805 |
(ALL xe::'A::type. IN xe x --> IN (xd xe) xa & xc (xd xe) = xb xe) --> |
|
8806 |
hollight.sum x xb = hollight.sum xa xc" |
|
8807 |
by (import hollight SUM_EQ_GENERAL) |
|
8808 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8809 |
definition CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list
|
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8810 |
=> 'q_57931 => 'q_57930 => 'q_57890" where |
| 17644 | 8811 |
"CASEWISE == |
| 19093 | 8812 |
SOME CASEWISE::(('q_57926::type => 'q_57930::type) *
|
8813 |
('q_57931::type
|
|
8814 |
=> 'q_57926::type => 'q_57890::type)) hollight.list |
|
8815 |
=> 'q_57931::type => 'q_57930::type => 'q_57890::type. |
|
8816 |
(ALL (f::'q_57931::type) x::'q_57930::type. |
|
8817 |
CASEWISE NIL f x = (SOME y::'q_57890::type. True)) & |
|
8818 |
(ALL (h::('q_57926::type => 'q_57930::type) *
|
|
8819 |
('q_57931::type => 'q_57926::type => 'q_57890::type))
|
|
8820 |
(t::(('q_57926::type => 'q_57930::type) *
|
|
8821 |
('q_57931::type
|
|
8822 |
=> 'q_57926::type => 'q_57890::type)) hollight.list) |
|
8823 |
(f::'q_57931::type) x::'q_57930::type. |
|
| 17644 | 8824 |
CASEWISE (CONS h t) f x = |
| 19093 | 8825 |
COND (EX y::'q_57926::type. fst h y = x) |
8826 |
(snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x))" |
|
| 17644 | 8827 |
|
8828 |
lemma DEF_CASEWISE: "CASEWISE = |
|
| 19093 | 8829 |
(SOME CASEWISE::(('q_57926::type => 'q_57930::type) *
|
8830 |
('q_57931::type
|
|
8831 |
=> 'q_57926::type => 'q_57890::type)) hollight.list |
|
8832 |
=> 'q_57931::type => 'q_57930::type => 'q_57890::type. |
|
8833 |
(ALL (f::'q_57931::type) x::'q_57930::type. |
|
8834 |
CASEWISE NIL f x = (SOME y::'q_57890::type. True)) & |
|
8835 |
(ALL (h::('q_57926::type => 'q_57930::type) *
|
|
8836 |
('q_57931::type => 'q_57926::type => 'q_57890::type))
|
|
8837 |
(t::(('q_57926::type => 'q_57930::type) *
|
|
8838 |
('q_57931::type
|
|
8839 |
=> 'q_57926::type => 'q_57890::type)) hollight.list) |
|
8840 |
(f::'q_57931::type) x::'q_57930::type. |
|
| 17644 | 8841 |
CASEWISE (CONS h t) f x = |
| 19093 | 8842 |
COND (EX y::'q_57926::type. fst h y = x) |
8843 |
(snd h f (SOME y::'q_57926::type. fst h y = x)) (CASEWISE t f x)))" |
|
| 17644 | 8844 |
by (import hollight DEF_CASEWISE) |
8845 |
||
8846 |
lemma CASEWISE: "(op &::bool => bool => bool) |
|
| 19093 | 8847 |
((op =::'q_57950::type => 'q_57950::type => bool) |
8848 |
((CASEWISE::(('q_57942::type => 'q_57990::type) *
|
|
8849 |
('q_57991::type
|
|
8850 |
=> 'q_57942::type => 'q_57950::type)) hollight.list |
|
8851 |
=> 'q_57991::type => 'q_57990::type => 'q_57950::type) |
|
8852 |
(NIL::(('q_57942::type => 'q_57990::type) *
|
|
8853 |
('q_57991::type
|
|
8854 |
=> 'q_57942::type => 'q_57950::type)) hollight.list) |
|
8855 |
(f::'q_57991::type) (x::'q_57990::type)) |
|
8856 |
((Eps::('q_57950::type => bool) => 'q_57950::type)
|
|
8857 |
(%y::'q_57950::type. True::bool))) |
|
8858 |
((op =::'q_57951::type => 'q_57951::type => bool) |
|
8859 |
((CASEWISE::(('q_57993::type => 'q_57990::type) *
|
|
8860 |
('q_57991::type
|
|
8861 |
=> 'q_57993::type => 'q_57951::type)) hollight.list |
|
8862 |
=> 'q_57991::type => 'q_57990::type => 'q_57951::type) |
|
8863 |
((CONS::('q_57993::type => 'q_57990::type) *
|
|
8864 |
('q_57991::type => 'q_57993::type => 'q_57951::type)
|
|
8865 |
=> (('q_57993::type => 'q_57990::type) *
|
|
8866 |
('q_57991::type
|
|
8867 |
=> 'q_57993::type => 'q_57951::type)) hollight.list |
|
8868 |
=> (('q_57993::type => 'q_57990::type) *
|
|
8869 |
('q_57991::type
|
|
8870 |
=> 'q_57993::type => 'q_57951::type)) hollight.list) |
|
8871 |
((Pair::('q_57993::type => 'q_57990::type)
|
|
8872 |
=> ('q_57991::type => 'q_57993::type => 'q_57951::type)
|
|
8873 |
=> ('q_57993::type => 'q_57990::type) *
|
|
8874 |
('q_57991::type => 'q_57993::type => 'q_57951::type))
|
|
8875 |
(s::'q_57993::type => 'q_57990::type) |
|
8876 |
(t::'q_57991::type => 'q_57993::type => 'q_57951::type)) |
|
8877 |
(clauses::(('q_57993::type => 'q_57990::type) *
|
|
8878 |
('q_57991::type
|
|
8879 |
=> 'q_57993::type => 'q_57951::type)) hollight.list)) |
|
| 17644 | 8880 |
f x) |
| 19093 | 8881 |
((COND::bool => 'q_57951::type => 'q_57951::type => 'q_57951::type) |
8882 |
((Ex::('q_57993::type => bool) => bool)
|
|
8883 |
(%y::'q_57993::type. |
|
8884 |
(op =::'q_57990::type => 'q_57990::type => bool) (s y) x)) |
|
8885 |
(t f ((Eps::('q_57993::type => bool) => 'q_57993::type)
|
|
8886 |
(%y::'q_57993::type. |
|
8887 |
(op =::'q_57990::type => 'q_57990::type => bool) (s y) x))) |
|
8888 |
((CASEWISE::(('q_57993::type => 'q_57990::type) *
|
|
8889 |
('q_57991::type
|
|
8890 |
=> 'q_57993::type => 'q_57951::type)) hollight.list |
|
8891 |
=> 'q_57991::type => 'q_57990::type => 'q_57951::type) |
|
| 17644 | 8892 |
clauses f x)))" |
8893 |
by (import hollight CASEWISE) |
|
8894 |
||
| 19093 | 8895 |
lemma CASEWISE_CASES: "ALL (clauses::(('q_58085::type => 'q_58082::type) *
|
8896 |
('q_58083::type
|
|
8897 |
=> 'q_58085::type => 'q_58092::type)) hollight.list) |
|
8898 |
(c::'q_58083::type) x::'q_58082::type. |
|
8899 |
(EX (s::'q_58085::type => 'q_58082::type) |
|
8900 |
(t::'q_58083::type => 'q_58085::type => 'q_58092::type) |
|
8901 |
a::'q_58085::type. |
|
| 17644 | 8902 |
MEM (s, t) clauses & s a = x & CASEWISE clauses c x = t c a) | |
| 19093 | 8903 |
~ (EX (s::'q_58085::type => 'q_58082::type) |
8904 |
(t::'q_58083::type => 'q_58085::type => 'q_58092::type) |
|
8905 |
a::'q_58085::type. MEM (s, t) clauses & s a = x) & |
|
8906 |
CASEWISE clauses c x = (SOME y::'q_58092::type. True)" |
|
| 17644 | 8907 |
by (import hollight CASEWISE_CASES) |
8908 |
||
8909 |
lemma CASEWISE_WORKS: "ALL (x::(('P::type => 'A::type) *
|
|
8910 |
('C::type => 'P::type => 'B::type)) hollight.list)
|
|
8911 |
xa::'C::type. |
|
8912 |
(ALL (s::'P::type => 'A::type) (t::'C::type => 'P::type => 'B::type) |
|
8913 |
(s'::'P::type => 'A::type) (t'::'C::type => 'P::type => 'B::type) |
|
8914 |
(xb::'P::type) y::'P::type. |
|
8915 |
MEM (s, t) x & MEM (s', t') x & s xb = s' y --> |
|
8916 |
t xa xb = t' xa y) --> |
|
8917 |
ALL_list |
|
8918 |
(GABS |
|
8919 |
(%f::('P::type => 'A::type) * ('C::type => 'P::type => 'B::type)
|
|
8920 |
=> bool. |
|
8921 |
ALL (s::'P::type => 'A::type) t::'C::type => 'P::type => 'B::type. |
|
8922 |
GEQ (f (s, t)) |
|
8923 |
(ALL xb::'P::type. CASEWISE x xa (s xb) = t xa xb))) |
|
8924 |
x" |
|
8925 |
by (import hollight CASEWISE_WORKS) |
|
8926 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8927 |
definition admissible :: "('q_58228 => 'q_58221 => bool)
|
| 19093 | 8928 |
=> (('q_58228 => 'q_58224) => 'q_58234 => bool)
|
8929 |
=> ('q_58234 => 'q_58221)
|
|
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8930 |
=> (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" where
|
| 17644 | 8931 |
"admissible == |
| 19093 | 8932 |
%(u::'q_58228::type => 'q_58221::type => bool) |
8933 |
(ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool)
|
|
8934 |
(ub::'q_58234::type => 'q_58221::type) |
|
8935 |
uc::('q_58228::type => 'q_58224::type)
|
|
8936 |
=> 'q_58234::type => 'q_58229::type. |
|
8937 |
ALL (f::'q_58228::type => 'q_58224::type) |
|
8938 |
(g::'q_58228::type => 'q_58224::type) a::'q_58234::type. |
|
| 17644 | 8939 |
ua f a & |
| 19093 | 8940 |
ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) --> |
| 17644 | 8941 |
uc f a = uc g a" |
8942 |
||
8943 |
lemma DEF_admissible: "admissible = |
|
| 19093 | 8944 |
(%(u::'q_58228::type => 'q_58221::type => bool) |
8945 |
(ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool)
|
|
8946 |
(ub::'q_58234::type => 'q_58221::type) |
|
8947 |
uc::('q_58228::type => 'q_58224::type)
|
|
8948 |
=> 'q_58234::type => 'q_58229::type. |
|
8949 |
ALL (f::'q_58228::type => 'q_58224::type) |
|
8950 |
(g::'q_58228::type => 'q_58224::type) a::'q_58234::type. |
|
| 17644 | 8951 |
ua f a & |
| 19093 | 8952 |
ua g a & (ALL z::'q_58228::type. u z (ub a) --> f z = g z) --> |
| 17644 | 8953 |
uc f a = uc g a)" |
8954 |
by (import hollight DEF_admissible) |
|
8955 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8956 |
definition tailadmissible :: "('A => 'A => bool)
|
| 17652 | 8957 |
=> (('A => 'B) => 'P => bool)
|
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8958 |
=> ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where
|
| 17644 | 8959 |
"tailadmissible == |
8960 |
%(u::'A::type => 'A::type => bool) |
|
8961 |
(ua::('A::type => 'B::type) => 'P::type => bool)
|
|
8962 |
(ub::'P::type => 'A::type) |
|
8963 |
uc::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
8964 |
EX (P::('A::type => 'B::type) => 'P::type => bool)
|
|
8965 |
(G::('A::type => 'B::type) => 'P::type => 'A::type)
|
|
8966 |
H::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
8967 |
(ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. |
|
8968 |
P f a & u y (G f a) --> u y (ub a)) & |
|
8969 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type. |
|
8970 |
(ALL z::'A::type. u z (ub a) --> f z = g z) --> |
|
8971 |
P f a = P g a & G f a = G g a & H f a = H g a) & |
|
8972 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
8973 |
ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a))" |
|
8974 |
||
8975 |
lemma DEF_tailadmissible: "tailadmissible = |
|
8976 |
(%(u::'A::type => 'A::type => bool) |
|
8977 |
(ua::('A::type => 'B::type) => 'P::type => bool)
|
|
8978 |
(ub::'P::type => 'A::type) |
|
8979 |
uc::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
8980 |
EX (P::('A::type => 'B::type) => 'P::type => bool)
|
|
8981 |
(G::('A::type => 'B::type) => 'P::type => 'A::type)
|
|
8982 |
H::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
8983 |
(ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. |
|
8984 |
P f a & u y (G f a) --> u y (ub a)) & |
|
8985 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) a::'P::type. |
|
8986 |
(ALL z::'A::type. u z (ub a) --> f z = g z) --> |
|
8987 |
P f a = P g a & G f a = G g a & H f a = H g a) & |
|
8988 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
8989 |
ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a)))" |
|
8990 |
by (import hollight DEF_tailadmissible) |
|
8991 |
||
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8992 |
definition superadmissible :: "('q_58378 => 'q_58378 => bool)
|
| 19093 | 8993 |
=> (('q_58378 => 'q_58380) => 'q_58386 => bool)
|
8994 |
=> ('q_58386 => 'q_58378)
|
|
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
34974
diff
changeset
|
8995 |
=> (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" where
|
| 17644 | 8996 |
"superadmissible == |
| 19093 | 8997 |
%(u::'q_58378::type => 'q_58378::type => bool) |
8998 |
(ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool)
|
|
8999 |
(ub::'q_58386::type => 'q_58378::type) |
|
9000 |
uc::('q_58378::type => 'q_58380::type)
|
|
9001 |
=> 'q_58386::type => 'q_58380::type. |
|
| 17644 | 9002 |
admissible u |
| 19093 | 9003 |
(%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub |
| 17644 | 9004 |
ua --> |
9005 |
tailadmissible u ua ub uc" |
|
9006 |
||
9007 |
lemma DEF_superadmissible: "superadmissible = |
|
| 19093 | 9008 |
(%(u::'q_58378::type => 'q_58378::type => bool) |
9009 |
(ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool)
|
|
9010 |
(ub::'q_58386::type => 'q_58378::type) |
|
9011 |
uc::('q_58378::type => 'q_58380::type)
|
|
9012 |
=> 'q_58386::type => 'q_58380::type. |
|
| 17644 | 9013 |
admissible u |
| 19093 | 9014 |
(%(f::'q_58378::type => 'q_58380::type) a::'q_58386::type. True) ub |
| 17644 | 9015 |
ua --> |
9016 |
tailadmissible u ua ub uc)" |
|
9017 |
by (import hollight DEF_superadmissible) |
|
9018 |
||
9019 |
lemma SUPERADMISSIBLE_COND: "ALL (x::'A::type => 'A::type => bool) |
|
9020 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9021 |
(xb::('A::type => 'B::type) => 'P::type => bool)
|
|
9022 |
(xc::'P::type => 'A::type) |
|
9023 |
(xd::('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9024 |
xe::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9025 |
admissible x xa xc xb & |
|
9026 |
superadmissible x |
|
9027 |
(%(f::'A::type => 'B::type) x::'P::type. xa f x & xb f x) xc xd & |
|
9028 |
superadmissible x |
|
9029 |
(%(f::'A::type => 'B::type) x::'P::type. xa f x & ~ xb f x) xc xe --> |
|
9030 |
superadmissible x xa xc |
|
9031 |
(%(f::'A::type => 'B::type) x::'P::type. |
|
9032 |
COND (xb f x) (xd f x) (xe f x))" |
|
9033 |
by (import hollight SUPERADMISSIBLE_COND) |
|
9034 |
||
9035 |
lemma ADMISSIBLE_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool) |
|
9036 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9037 |
(xb::'P::type => 'A::type) |
|
9038 |
xc::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9039 |
admissible x xa xb xc --> superadmissible x xa xb xc" |
|
9040 |
by (import hollight ADMISSIBLE_IMP_SUPERADMISSIBLE) |
|
9041 |
||
9042 |
lemma TAIL_IMP_SUPERADMISSIBLE: "ALL (x::'A::type => 'A::type => bool) |
|
9043 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9044 |
(xb::'P::type => 'A::type) |
|
9045 |
xc::('A::type => 'B::type) => 'P::type => 'A::type.
|
|
9046 |
admissible x xa xb xc & |
|
9047 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
9048 |
xa f a --> (ALL y::'A::type. x y (xc f a) --> x y (xb a))) --> |
|
9049 |
superadmissible x xa xb |
|
9050 |
(%(f::'A::type => 'B::type) x::'P::type. f (xc f x))" |
|
9051 |
by (import hollight TAIL_IMP_SUPERADMISSIBLE) |
|
9052 |
||
| 19093 | 9053 |
lemma ADMISSIBLE_COND: "ALL (u_353::'A::type => 'q_58766::type => bool) |
| 17644 | 9054 |
(p::('A::type => 'B::type) => 'P::type => bool)
|
9055 |
(P::('A::type => 'B::type) => 'P::type => bool)
|
|
| 19093 | 9056 |
(s::'P::type => 'q_58766::type) |
| 17644 | 9057 |
(h::('A::type => 'B::type) => 'P::type => 'B::type)
|
9058 |
k::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
| 19093 | 9059 |
admissible u_353 p s P & |
9060 |
admissible u_353 (%(f::'A::type => 'B::type) x::'P::type. p f x & P f x) |
|
| 17644 | 9061 |
s h & |
| 19093 | 9062 |
admissible u_353 |
| 17644 | 9063 |
(%(f::'A::type => 'B::type) x::'P::type. p f x & ~ P f x) s k --> |
| 19093 | 9064 |
admissible u_353 p s |
| 17644 | 9065 |
(%(f::'A::type => 'B::type) x::'P::type. COND (P f x) (h f x) (k f x))" |
9066 |
by (import hollight ADMISSIBLE_COND) |
|
9067 |
||
| 19093 | 9068 |
lemma ADMISSIBLE_CONST: "admissible (u_353::'q_58841::type => 'q_58840::type => bool) |
9069 |
(p::('q_58841::type => 'q_58842::type) => 'q_58843::type => bool)
|
|
9070 |
(s::'q_58843::type => 'q_58840::type) |
|
9071 |
(%f::'q_58841::type => 'q_58842::type. c::'q_58843::type => 'q_58844::type)" |
|
| 17644 | 9072 |
by (import hollight ADMISSIBLE_CONST) |
9073 |
||
9074 |
lemma ADMISSIBLE_COMB: "ALL (x::'A::type => 'A::type => bool) |
|
9075 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9076 |
(xb::'P::type => 'A::type) |
|
9077 |
(xc::('A::type => 'B::type) => 'P::type => 'C::type => 'D::type)
|
|
9078 |
xd::('A::type => 'B::type) => 'P::type => 'C::type.
|
|
9079 |
admissible x xa xb xc & admissible x xa xb xd --> |
|
9080 |
admissible x xa xb |
|
9081 |
(%(f::'A::type => 'B::type) x::'P::type. xc f x (xd f x))" |
|
9082 |
by (import hollight ADMISSIBLE_COMB) |
|
9083 |
||
9084 |
lemma ADMISSIBLE_BASE: "ALL (x::'A::type => 'A::type => bool) |
|
9085 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9086 |
(xb::'P::type => 'A::type) xc::'P::type => 'A::type. |
|
9087 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
9088 |
xa f a --> x (xc a) (xb a)) --> |
|
9089 |
admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc x))" |
|
9090 |
by (import hollight ADMISSIBLE_BASE) |
|
9091 |
||
9092 |
lemma ADMISSIBLE_NEST: "ALL (x::'A::type => 'A::type => bool) |
|
9093 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9094 |
(xb::'P::type => 'A::type) |
|
9095 |
xc::('A::type => 'B::type) => 'P::type => 'A::type.
|
|
9096 |
admissible x xa xb xc & |
|
9097 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
9098 |
xa f a --> x (xc f a) (xb a)) --> |
|
9099 |
admissible x xa xb (%(f::'A::type => 'B::type) x::'P::type. f (xc f x))" |
|
9100 |
by (import hollight ADMISSIBLE_NEST) |
|
9101 |
||
9102 |
lemma ADMISSIBLE_LAMBDA: "ALL (x::'A::type => 'A::type => bool) |
|
9103 |
(xa::('A::type => 'B::type) => 'P::type => bool)
|
|
9104 |
(xb::'P::type => 'A::type) |
|
9105 |
xc::'C::type => ('A::type => 'B::type) => 'P::type => 'D::type.
|
|
9106 |
(ALL xd::'C::type. admissible x xa xb (xc xd)) --> |
|
9107 |
admissible x xa xb |
|
9108 |
(%(f::'A::type => 'B::type) (x::'P::type) u::'C::type. xc u f x)" |
|
9109 |
by (import hollight ADMISSIBLE_LAMBDA) |
|
9110 |
||
| 19093 | 9111 |
lemma ADMISSIBLE_NSUM: "ALL (x::'B::type => 'A::type => bool) |
9112 |
(xa::('B::type => 'C::type) => 'P::type => bool)
|
|
9113 |
(xb::'P::type => 'A::type) |
|
9114 |
(xc::('B::type => 'C::type) => 'P::type => nat => nat)
|
|
9115 |
(xd::'P::type => nat) xe::'P::type => nat. |
|
9116 |
(ALL xf::nat. |
|
9117 |
admissible x |
|
9118 |
(%(f::'B::type => 'C::type) x::'P::type. |
|
9119 |
<= (xd x) xf & <= xf (xe x) & xa f x) |
|
9120 |
xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) --> |
|
9121 |
admissible x xa xb |
|
9122 |
(%(f::'B::type => 'C::type) x::'P::type. |
|
9123 |
nsum (dotdot (xd x) (xe x)) (xc f x))" |
|
9124 |
by (import hollight ADMISSIBLE_NSUM) |
|
9125 |
||
9126 |
lemma ADMISSIBLE_SUM: "ALL (x::'B::type => 'A::type => bool) |
|
9127 |
(xa::('B::type => 'C::type) => 'P::type => bool)
|
|
9128 |
(xb::'P::type => 'A::type) |
|
9129 |
(xc::('B::type => 'C::type) => 'P::type => nat => hollight.real)
|
|
9130 |
(xd::'P::type => nat) xe::'P::type => nat. |
|
9131 |
(ALL xf::nat. |
|
9132 |
admissible x |
|
9133 |
(%(f::'B::type => 'C::type) x::'P::type. |
|
9134 |
<= (xd x) xf & <= xf (xe x) & xa f x) |
|
9135 |
xb (%(f::'B::type => 'C::type) x::'P::type. xc f x xf)) --> |
|
9136 |
admissible x xa xb |
|
9137 |
(%(f::'B::type => 'C::type) x::'P::type. |
|
9138 |
hollight.sum (dotdot (xd x) (xe x)) (xc f x))" |
|
9139 |
by (import hollight ADMISSIBLE_SUM) |
|
9140 |
||
9141 |
lemma WF_REC_CASES: "ALL (u_353::'A::type => 'A::type => bool) |
|
| 17644 | 9142 |
clauses::(('P::type => 'A::type) *
|
9143 |
(('A::type => 'B::type)
|
|
9144 |
=> 'P::type => 'B::type)) hollight.list. |
|
| 19093 | 9145 |
WF u_353 & |
| 17644 | 9146 |
ALL_list |
9147 |
(GABS |
|
9148 |
(%f::('P::type => 'A::type) *
|
|
9149 |
(('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9150 |
=> bool. |
|
9151 |
ALL (s::'P::type => 'A::type) |
|
9152 |
t::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9153 |
GEQ (f (s, t)) |
|
9154 |
(EX (P::('A::type => 'B::type) => 'P::type => bool)
|
|
9155 |
(G::('A::type => 'B::type) => 'P::type => 'A::type)
|
|
9156 |
H::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9157 |
(ALL (f::'A::type => 'B::type) (a::'P::type) y::'A::type. |
|
| 19093 | 9158 |
P f a & u_353 y (G f a) --> u_353 y (s a)) & |
| 17644 | 9159 |
(ALL (f::'A::type => 'B::type) (g::'A::type => 'B::type) |
9160 |
a::'P::type. |
|
| 19093 | 9161 |
(ALL z::'A::type. u_353 z (s a) --> f z = g z) --> |
| 17644 | 9162 |
P f a = P g a & G f a = G g a & H f a = H g a) & |
9163 |
(ALL (f::'A::type => 'B::type) a::'P::type. |
|
9164 |
t f a = COND (P f a) (f (G f a)) (H f a))))) |
|
9165 |
clauses --> |
|
9166 |
(EX f::'A::type => 'B::type. ALL x::'A::type. f x = CASEWISE clauses f x)" |
|
9167 |
by (import hollight WF_REC_CASES) |
|
9168 |
||
9169 |
lemma RECURSION_CASEWISE: "ALL clauses::(('P::type => 'A::type) *
|
|
9170 |
(('A::type => 'B::type)
|
|
9171 |
=> 'P::type => 'B::type)) hollight.list. |
|
9172 |
(EX u::'A::type => 'A::type => bool. |
|
9173 |
WF u & |
|
9174 |
ALL_list |
|
9175 |
(GABS |
|
9176 |
(%f::('P::type => 'A::type) *
|
|
9177 |
(('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9178 |
=> bool. |
|
9179 |
ALL (s::'P::type => 'A::type) |
|
9180 |
t::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9181 |
GEQ (f (s, t)) |
|
9182 |
(tailadmissible u |
|
9183 |
(%(f::'A::type => 'B::type) a::'P::type. True) s t))) |
|
9184 |
clauses) & |
|
9185 |
(ALL (x::'P::type => 'A::type) |
|
9186 |
(xa::('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9187 |
(xb::'P::type => 'A::type) |
|
9188 |
(xc::('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9189 |
(xd::'A::type => 'B::type) (xe::'P::type) xf::'P::type. |
|
9190 |
MEM (x, xa) clauses & MEM (xb, xc) clauses --> |
|
9191 |
x xe = xb xf --> xa xd xe = xc xd xf) --> |
|
9192 |
(EX f::'A::type => 'B::type. |
|
9193 |
ALL_list |
|
9194 |
(GABS |
|
9195 |
(%fa::('P::type => 'A::type) *
|
|
9196 |
(('A::type => 'B::type) => 'P::type => 'B::type)
|
|
9197 |
=> bool. |
|
9198 |
ALL (s::'P::type => 'A::type) |
|
9199 |
t::('A::type => 'B::type) => 'P::type => 'B::type.
|
|
9200 |
GEQ (fa (s, t)) (ALL x::'P::type. f (s x) = t f x))) |
|
9201 |
clauses)" |
|
9202 |
by (import hollight RECURSION_CASEWISE) |
|
9203 |
||
| 19093 | 9204 |
lemma cth: "ALL (p1::'A::type => 'q_59947::type) |
9205 |
(p2::'q_59958::type => 'A::type => 'q_59952::type) |
|
9206 |
(p1'::'A::type => 'q_59947::type) |
|
9207 |
p2'::'q_59958::type => 'A::type => 'q_59952::type. |
|
9208 |
(ALL (c::'q_59958::type) (x::'A::type) y::'A::type. |
|
| 17644 | 9209 |
p1 x = p1' y --> p2 c x = p2' c y) --> |
| 19093 | 9210 |
(ALL (c::'q_59958::type) (x::'A::type) y::'A::type. |
| 17644 | 9211 |
p1' x = p1 y --> p2' c x = p2 c y)" |
9212 |
by (import hollight cth) |
|
9213 |
||
| 19093 | 9214 |
lemma RECURSION_CASEWISE_PAIRWISE: "ALL x::(('q_59995::type => 'q_59975::type) *
|
9215 |
(('q_59975::type => 'q_59991::type)
|
|
9216 |
=> 'q_59995::type => 'q_59991::type)) hollight.list. |
|
9217 |
(EX u::'q_59975::type => 'q_59975::type => bool. |
|
| 17644 | 9218 |
WF u & |
9219 |
ALL_list |
|
9220 |
(GABS |
|
| 19093 | 9221 |
(%f::('q_59995::type => 'q_59975::type) *
|
9222 |
(('q_59975::type => 'q_59991::type)
|
|
9223 |
=> 'q_59995::type => 'q_59991::type) |
|
| 17644 | 9224 |
=> bool. |
| 19093 | 9225 |
ALL (s::'q_59995::type => 'q_59975::type) |
9226 |
t::('q_59975::type => 'q_59991::type)
|
|
9227 |
=> 'q_59995::type => 'q_59991::type. |
|
| 17644 | 9228 |
GEQ (f (s, t)) |
9229 |
(tailadmissible u |
|
| 19093 | 9230 |
(%(f::'q_59975::type => 'q_59991::type) |
9231 |
a::'q_59995::type. True) |
|
| 17644 | 9232 |
s t))) |
9233 |
x) & |
|
9234 |
ALL_list |
|
9235 |
(GABS |
|
| 19093 | 9236 |
(%f::('q_59995::type => 'q_59975::type) *
|
9237 |
(('q_59975::type => 'q_59991::type)
|
|
9238 |
=> 'q_59995::type => 'q_59991::type) |
|
| 17644 | 9239 |
=> bool. |
| 19093 | 9240 |
ALL (a::'q_59995::type => 'q_59975::type) |
9241 |
b::('q_59975::type => 'q_59991::type)
|
|
9242 |
=> 'q_59995::type => 'q_59991::type. |
|
| 17644 | 9243 |
GEQ (f (a, b)) |
| 19093 | 9244 |
(ALL (c::'q_59975::type => 'q_59991::type) (x::'q_59995::type) |
9245 |
y::'q_59995::type. a x = a y --> b c x = b c y))) |
|
| 17644 | 9246 |
x & |
9247 |
PAIRWISE |
|
9248 |
(GABS |
|
| 19093 | 9249 |
(%f::('q_59995::type => 'q_59975::type) *
|
9250 |
(('q_59975::type => 'q_59991::type)
|
|
9251 |
=> 'q_59995::type => 'q_59991::type) |
|
9252 |
=> ('q_59995::type => 'q_59975::type) *
|
|
9253 |
(('q_59975::type => 'q_59991::type)
|
|
9254 |
=> 'q_59995::type => 'q_59991::type) |
|
| 17644 | 9255 |
=> bool. |
| 19093 | 9256 |
ALL (s::'q_59995::type => 'q_59975::type) |
9257 |
t::('q_59975::type => 'q_59991::type)
|
|
9258 |
=> 'q_59995::type => 'q_59991::type. |
|
| 17644 | 9259 |
GEQ (f (s, t)) |
9260 |
(GABS |
|
| 19093 | 9261 |
(%f::('q_59995::type => 'q_59975::type) *
|
9262 |
(('q_59975::type => 'q_59991::type)
|
|
9263 |
=> 'q_59995::type => 'q_59991::type) |
|
| 17644 | 9264 |
=> bool. |
| 19093 | 9265 |
ALL (s'::'q_59995::type => 'q_59975::type) |
9266 |
t'::('q_59975::type => 'q_59991::type)
|
|
9267 |
=> 'q_59995::type => 'q_59991::type. |
|
| 17644 | 9268 |
GEQ (f (s', t')) |
| 19093 | 9269 |
(ALL (c::'q_59975::type => 'q_59991::type) |
9270 |
(x::'q_59995::type) y::'q_59995::type. |
|
| 17644 | 9271 |
s x = s' y --> t c x = t' c y))))) |
9272 |
x --> |
|
| 19093 | 9273 |
(EX f::'q_59975::type => 'q_59991::type. |
| 17644 | 9274 |
ALL_list |
9275 |
(GABS |
|
| 19093 | 9276 |
(%fa::('q_59995::type => 'q_59975::type) *
|
9277 |
(('q_59975::type => 'q_59991::type)
|
|
9278 |
=> 'q_59995::type => 'q_59991::type) |
|
| 17644 | 9279 |
=> bool. |
| 19093 | 9280 |
ALL (s::'q_59995::type => 'q_59975::type) |
9281 |
t::('q_59975::type => 'q_59991::type)
|
|
9282 |
=> 'q_59995::type => 'q_59991::type. |
|
9283 |
GEQ (fa (s, t)) (ALL x::'q_59995::type. f (s x) = t f x))) |
|
| 17644 | 9284 |
x)" |
9285 |
by (import hollight RECURSION_CASEWISE_PAIRWISE) |
|
9286 |
||
| 19093 | 9287 |
lemma SUPERADMISSIBLE_T: "superadmissible (u_353::'q_60105::type => 'q_60105::type => bool) |
9288 |
(%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True) |
|
9289 |
(s::'q_60111::type => 'q_60105::type) |
|
9290 |
(t::('q_60105::type => 'q_60107::type)
|
|
9291 |
=> 'q_60111::type => 'q_60107::type) = |
|
9292 |
tailadmissible u_353 |
|
9293 |
(%(f::'q_60105::type => 'q_60107::type) x::'q_60111::type. True) s t" |
|
| 17644 | 9294 |
by (import hollight SUPERADMISSIBLE_T) |
9295 |
||
9296 |
;end_setup |
|
9297 |
||
9298 |
end |
|
9299 |