separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
2 (* Authors: Jeremy Avigad and Amine Chaieb *)
4 section {* Generic transfer machinery; specific transfer from nats to ints and back. *}
10 subsection {* Generic transfer machinery *}
12 definition transfer_morphism:: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
13 where "transfer_morphism f A \<longleftrightarrow> True"
15 lemma transfer_morphismI[intro]: "transfer_morphism f A"
16 by (simp add: transfer_morphism_def)
18 ML_file "Tools/legacy_transfer.ML"
21 subsection {* Set up transfer from nat to int *}
23 text {* set up transfer direction *}
25 lemma transfer_morphism_nat_int: "transfer_morphism nat (op <= (0::int))" ..
27 declare transfer_morphism_nat_int [transfer add
33 text {* basic functions and relations *}
35 lemma transfer_nat_int_numerals [transfer key: transfer_morphism_nat_int]:
43 tsub :: "int \<Rightarrow> int \<Rightarrow> int"
45 "tsub x y = (if x >= y then x - y else 0)"
47 lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y"
48 by (simp add: tsub_def)
50 lemma transfer_nat_int_functions [transfer key: transfer_morphism_nat_int]:
51 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)"
52 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)"
53 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)"
54 "(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)"
55 by (auto simp add: eq_nat_nat_iff nat_mult_distrib
56 nat_power_eq tsub_def)
58 lemma transfer_nat_int_function_closures [transfer key: transfer_morphism_nat_int]:
59 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0"
60 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0"
61 "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0"
62 "(x::int) >= 0 \<Longrightarrow> x^n >= 0"
68 by (auto simp add: zero_le_mult_iff tsub_def)
70 lemma transfer_nat_int_relations [transfer key: transfer_morphism_nat_int]:
71 "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
72 (nat (x::int) = nat y) = (x = y)"
73 "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
74 (nat (x::int) < nat y) = (x < y)"
75 "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
76 (nat (x::int) <= nat y) = (x <= y)"
77 "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow>
78 (nat (x::int) dvd nat y) = (x dvd y)"
79 by (auto simp add: zdvd_int)
82 text {* first-order quantifiers *}
84 lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"
85 by (simp split add: split_nat)
87 lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))"
89 assume "\<exists>x. P x"
90 then obtain x where "P x" ..
91 then have "int x \<ge> 0 \<and> P (nat (int x))" by simp
92 then show "\<exists>x\<ge>0. P (nat x)" ..
94 assume "\<exists>x\<ge>0. P (nat x)"
95 then show "\<exists>x. P x" by auto
98 lemma transfer_nat_int_quantifiers [transfer key: transfer_morphism_nat_int]:
99 "(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))"
100 "(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))"
101 by (rule all_nat, rule ex_nat)
103 (* should we restrict these? *)
104 lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
105 (ALL x. Q x \<longrightarrow> P x) = (ALL x. Q x \<longrightarrow> P' x)"
108 lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow>
109 (EX x. Q x \<and> P x) = (EX x. Q x \<and> P' x)"
112 declare transfer_morphism_nat_int [transfer add
113 cong: all_cong ex_cong]
118 lemma nat_if_cong [transfer key: transfer_morphism_nat_int]:
119 "(if P then (nat x) else (nat y)) = nat (if P then x else y)"
123 text {* operations with sets *}
126 nat_set :: "int set \<Rightarrow> bool"
128 "nat_set S = (ALL x:S. x >= 0)"
130 lemma transfer_nat_int_set_functions:
131 "card A = card (int ` A)"
132 "{} = nat ` ({}::int set)"
133 "A Un B = nat ` (int ` A Un int ` B)"
134 "A Int B = nat ` (int ` A Int int ` B)"
135 "{x. P x} = nat ` {x. x >= 0 & P(nat x)}"
136 apply (rule card_image [symmetric])
137 apply (auto simp add: inj_on_def image_def)
138 apply (rule_tac x = "int x" in bexI)
140 apply (rule_tac x = "int x" in bexI)
142 apply (rule_tac x = "int x" in bexI)
144 apply (rule_tac x = "int x" in exI)
148 lemma transfer_nat_int_set_function_closures:
150 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
151 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
152 "nat_set {x. x >= 0 & P x}"
154 "nat_set A \<Longrightarrow> x : A \<Longrightarrow> x >= 0" (* does it hurt to turn this on? *)
155 unfolding nat_set_def apply auto
158 lemma transfer_nat_int_set_relations:
159 "(finite A) = (finite (int ` A))"
160 "(x : A) = (int x : int ` A)"
161 "(A = B) = (int ` A = int ` B)"
162 "(A < B) = (int ` A < int ` B)"
163 "(A <= B) = (int ` A <= int ` B)"
165 apply (erule finite_imageI)
166 apply (erule finite_imageD)
167 apply (auto simp add: image_def set_eq_iff inj_on_def)
168 apply (drule_tac x = "int x" in spec, auto)
169 apply (drule_tac x = "int x" in spec, auto)
170 apply (drule_tac x = "int x" in spec, auto)
173 lemma transfer_nat_int_set_return_embed: "nat_set A \<Longrightarrow>
175 by (auto simp add: nat_set_def image_def)
177 lemma transfer_nat_int_set_cong: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow>
178 {(x::int). x >= 0 & P x} = {x. x >= 0 & P' x}"
181 declare transfer_morphism_nat_int [transfer add
182 return: transfer_nat_int_set_functions
183 transfer_nat_int_set_function_closures
184 transfer_nat_int_set_relations
185 transfer_nat_int_set_return_embed
186 cong: transfer_nat_int_set_cong
190 text {* setsum and setprod *}
192 (* this handles the case where the *domain* of f is nat *)
193 lemma transfer_nat_int_sum_prod:
194 "setsum f A = setsum (%x. f (nat x)) (int ` A)"
195 "setprod f A = setprod (%x. f (nat x)) (int ` A)"
196 apply (subst setsum.reindex)
197 apply (unfold inj_on_def, auto)
198 apply (subst setprod.reindex)
199 apply (unfold inj_on_def o_def, auto)
202 (* this handles the case where the *range* of f is nat *)
203 lemma transfer_nat_int_sum_prod2:
204 "setsum f A = nat(setsum (%x. int (f x)) A)"
205 "setprod f A = nat(setprod (%x. int (f x)) A)"
206 apply (subst int_setsum [symmetric])
208 apply (subst int_setprod [symmetric])
212 lemma transfer_nat_int_sum_prod_closure:
213 "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
214 "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
215 unfolding nat_set_def
216 apply (rule setsum_nonneg)
218 apply (rule setprod_nonneg)
222 (* this version doesn't work, even with nat_set A \<Longrightarrow>
223 x : A \<Longrightarrow> x >= 0 turned on. Why not?
225 also: what does =simp=> do?
227 lemma transfer_nat_int_sum_prod_closure:
228 "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
229 "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
230 unfolding nat_set_def simp_implies_def
231 apply (rule setsum_nonneg)
233 apply (rule setprod_nonneg)
238 (* Making A = B in this lemma doesn't work. Why not?
239 Also, why aren't setsum.cong and setprod.cong enough,
240 with the previously mentioned rule turned on? *)
242 lemma transfer_nat_int_sum_prod_cong:
243 "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
244 setsum f A = setsum g B"
245 "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
246 setprod f A = setprod g B"
247 unfolding nat_set_def
248 apply (subst setsum.cong, assumption)
250 apply (subst setprod.cong, assumption, auto)
253 declare transfer_morphism_nat_int [transfer add
254 return: transfer_nat_int_sum_prod transfer_nat_int_sum_prod2
255 transfer_nat_int_sum_prod_closure
256 cong: transfer_nat_int_sum_prod_cong]
259 subsection {* Set up transfer from int to nat *}
261 text {* set up transfer direction *}
263 lemma transfer_morphism_int_nat: "transfer_morphism int (\<lambda>n. True)" ..
265 declare transfer_morphism_int_nat [transfer add
272 text {* basic functions and relations *}
275 is_nat :: "int \<Rightarrow> bool"
277 "is_nat x = (x >= 0)"
279 lemma transfer_int_nat_numerals:
286 lemma transfer_int_nat_functions:
287 "(int x) + (int y) = int (x + y)"
288 "(int x) * (int y) = int (x * y)"
289 "tsub (int x) (int y) = int (x - y)"
290 "(int x)^n = int (x^n)"
291 by (auto simp add: int_mult tsub_def int_power)
293 lemma transfer_int_nat_function_closures:
294 "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)"
295 "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)"
296 "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)"
297 "is_nat x \<Longrightarrow> is_nat (x^n)"
303 by (simp_all only: is_nat_def transfer_nat_int_function_closures)
305 lemma transfer_int_nat_relations:
306 "(int x = int y) = (x = y)"
307 "(int x < int y) = (x < y)"
308 "(int x <= int y) = (x <= y)"
309 "(int x dvd int y) = (x dvd y)"
310 by (auto simp add: zdvd_int)
312 declare transfer_morphism_int_nat [transfer add return:
313 transfer_int_nat_numerals
314 transfer_int_nat_functions
315 transfer_int_nat_function_closures
316 transfer_int_nat_relations
320 text {* first-order quantifiers *}
322 lemma transfer_int_nat_quantifiers:
323 "(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))"
324 "(EX (x::int) >= 0. P x) = (EX (x::nat). P (int x))"
325 apply (subst all_nat)
331 declare transfer_morphism_int_nat [transfer add
332 return: transfer_int_nat_quantifiers]
337 lemma int_if_cong: "(if P then (int x) else (int y)) =
338 int (if P then x else y)"
341 declare transfer_morphism_int_nat [transfer add return: int_if_cong]
345 text {* operations with sets *}
347 lemma transfer_int_nat_set_functions:
348 "nat_set A \<Longrightarrow> card A = card (nat ` A)"
349 "{} = int ` ({}::nat set)"
350 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)"
351 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Int B = int ` (nat ` A Int nat ` B)"
352 "{x. x >= 0 & P x} = int ` {x. P(int x)}"
353 (* need all variants of these! *)
354 by (simp_all only: is_nat_def transfer_nat_int_set_functions
355 transfer_nat_int_set_function_closures
356 transfer_nat_int_set_return_embed nat_0_le
357 cong: transfer_nat_int_set_cong)
359 lemma transfer_int_nat_set_function_closures:
361 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
362 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
363 "nat_set {x. x >= 0 & P x}"
365 "nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x"
366 by (simp_all only: transfer_nat_int_set_function_closures is_nat_def)
368 lemma transfer_int_nat_set_relations:
369 "nat_set A \<Longrightarrow> finite A = finite (nat ` A)"
370 "is_nat x \<Longrightarrow> nat_set A \<Longrightarrow> (x : A) = (nat x : nat ` A)"
371 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A = B) = (nat ` A = nat ` B)"
372 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A < B) = (nat ` A < nat ` B)"
373 "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A <= B) = (nat ` A <= nat ` B)"
374 by (simp_all only: is_nat_def transfer_nat_int_set_relations
375 transfer_nat_int_set_return_embed nat_0_le)
377 lemma transfer_int_nat_set_return_embed: "nat ` int ` A = A"
378 by (simp only: transfer_nat_int_set_relations
379 transfer_nat_int_set_function_closures
380 transfer_nat_int_set_return_embed nat_0_le)
382 lemma transfer_int_nat_set_cong: "(!!x. P x = P' x) \<Longrightarrow>
383 {(x::nat). P x} = {x. P' x}"
386 declare transfer_morphism_int_nat [transfer add
387 return: transfer_int_nat_set_functions
388 transfer_int_nat_set_function_closures
389 transfer_int_nat_set_relations
390 transfer_int_nat_set_return_embed
391 cong: transfer_int_nat_set_cong
395 text {* setsum and setprod *}
397 (* this handles the case where the *domain* of f is int *)
398 lemma transfer_int_nat_sum_prod:
399 "nat_set A \<Longrightarrow> setsum f A = setsum (%x. f (int x)) (nat ` A)"
400 "nat_set A \<Longrightarrow> setprod f A = setprod (%x. f (int x)) (nat ` A)"
401 apply (subst setsum.reindex)
402 apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff)
403 apply (subst setprod.reindex)
404 apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff
408 (* this handles the case where the *range* of f is int *)
409 lemma transfer_int_nat_sum_prod2:
410 "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> setsum f A = int(setsum (%x. nat (f x)) A)"
411 "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow>
412 setprod f A = int(setprod (%x. nat (f x)) A)"
414 apply (subst int_setsum, auto)
415 apply (subst int_setprod, auto simp add: cong: setprod.cong)
418 declare transfer_morphism_int_nat [transfer add
419 return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2
420 cong: setsum.cong setprod.cong]