| author | nipkow | 
| Tue, 20 Oct 2009 14:44:02 +0200 | |
| changeset 33019 | bcf56a64ce1a | 
| parent 32558 | e6e1fc2e73cb | 
| child 33318 | ddd97d9dfbfb | 
| permissions | -rw-r--r-- | 
| 31708 | 1 | |
| 32554 | 2 | (* Authors: Jeremy Avigad and Amine Chaieb *) | 
| 31708 | 3 | |
| 32554 | 4 | header {* Sets up transfer from nats to ints and back. *}
 | 
| 31708 | 5 | |
| 32558 | 6 | theory Nat_Transfer | 
| 31708 | 7 | imports Main Parity | 
| 8 | begin | |
| 9 | ||
| 10 | subsection {* Set up transfer from nat to int *}
 | |
| 11 | ||
| 12 | (* set up transfer direction *) | |
| 13 | ||
| 14 | lemma TransferMorphism_nat_int: "TransferMorphism nat (op <= (0::int))" | |
| 15 | by (simp add: TransferMorphism_def) | |
| 16 | ||
| 17 | declare TransferMorphism_nat_int[transfer | |
| 18 | add mode: manual | |
| 19 | return: nat_0_le | |
| 20 | labels: natint | |
| 21 | ] | |
| 22 | ||
| 23 | (* basic functions and relations *) | |
| 24 | ||
| 25 | lemma transfer_nat_int_numerals: | |
| 26 | "(0::nat) = nat 0" | |
| 27 | "(1::nat) = nat 1" | |
| 28 | "(2::nat) = nat 2" | |
| 29 | "(3::nat) = nat 3" | |
| 30 | by auto | |
| 31 | ||
| 32 | definition | |
| 33 | tsub :: "int \<Rightarrow> int \<Rightarrow> int" | |
| 34 | where | |
| 35 | "tsub x y = (if x >= y then x - y else 0)" | |
| 36 | ||
| 37 | lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y" | |
| 38 | by (simp add: tsub_def) | |
| 39 | ||
| 40 | ||
| 41 | lemma transfer_nat_int_functions: | |
| 42 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)" | |
| 43 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)" | |
| 44 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)" | |
| 45 | "(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)" | |
| 46 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)" | |
| 47 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)" | |
| 48 | by (auto simp add: eq_nat_nat_iff nat_mult_distrib | |
| 49 | nat_power_eq nat_div_distrib nat_mod_distrib tsub_def) | |
| 50 | ||
| 51 | lemma transfer_nat_int_function_closures: | |
| 52 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0" | |
| 53 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0" | |
| 54 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0" | |
| 55 | "(x::int) >= 0 \<Longrightarrow> x^n >= 0" | |
| 56 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0" | |
| 57 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0" | |
| 58 | "(0::int) >= 0" | |
| 59 | "(1::int) >= 0" | |
| 60 | "(2::int) >= 0" | |
| 61 | "(3::int) >= 0" | |
| 62 | "int z >= 0" | |
| 63 | apply (auto simp add: zero_le_mult_iff tsub_def) | |
| 64 | apply (case_tac "y = 0") | |
| 65 | apply auto | |
| 66 | apply (subst pos_imp_zdiv_nonneg_iff, auto) | |
| 67 | apply (case_tac "y = 0") | |
| 68 | apply force | |
| 69 | apply (rule pos_mod_sign) | |
| 70 | apply arith | |
| 71 | done | |
| 72 | ||
| 73 | lemma transfer_nat_int_relations: | |
| 74 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | |
| 75 | (nat (x::int) = nat y) = (x = y)" | |
| 76 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | |
| 77 | (nat (x::int) < nat y) = (x < y)" | |
| 78 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | |
| 79 | (nat (x::int) <= nat y) = (x <= y)" | |
| 80 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | |
| 81 | (nat (x::int) dvd nat y) = (x dvd y)" | |
| 32558 | 82 | by (auto simp add: zdvd_int) | 
| 31708 | 83 | |
| 84 | declare TransferMorphism_nat_int[transfer add return: | |
| 85 | transfer_nat_int_numerals | |
| 86 | transfer_nat_int_functions | |
| 87 | transfer_nat_int_function_closures | |
| 88 | transfer_nat_int_relations | |
| 89 | ] | |
| 90 | ||
| 91 | ||
| 92 | (* first-order quantifiers *) | |
| 93 | ||
| 94 | lemma transfer_nat_int_quantifiers: | |
| 95 | "(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))" | |
| 96 | "(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))" | |
| 97 | by (rule all_nat, rule ex_nat) | |
| 98 | ||
| 99 | (* should we restrict these? *) | |
| 100 | lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 101 | (ALL x. Q x \<longrightarrow> P x) = (ALL x. Q x \<longrightarrow> P' x)" | |
| 102 | by auto | |
| 103 | ||
| 104 | lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 105 | (EX x. Q x \<and> P x) = (EX x. Q x \<and> P' x)" | |
| 106 | by auto | |
| 107 | ||
| 108 | declare TransferMorphism_nat_int[transfer add | |
| 109 | return: transfer_nat_int_quantifiers | |
| 110 | cong: all_cong ex_cong] | |
| 111 | ||
| 112 | ||
| 113 | (* if *) | |
| 114 | ||
| 115 | lemma nat_if_cong: "(if P then (nat x) else (nat y)) = | |
| 116 | nat (if P then x else y)" | |
| 117 | by auto | |
| 118 | ||
| 119 | declare TransferMorphism_nat_int [transfer add return: nat_if_cong] | |
| 120 | ||
| 121 | ||
| 122 | (* operations with sets *) | |
| 123 | ||
| 124 | definition | |
| 125 | nat_set :: "int set \<Rightarrow> bool" | |
| 126 | where | |
| 127 | "nat_set S = (ALL x:S. x >= 0)" | |
| 128 | ||
| 129 | lemma transfer_nat_int_set_functions: | |
| 130 | "card A = card (int ` A)" | |
| 131 |     "{} = nat ` ({}::int set)"
 | |
| 132 | "A Un B = nat ` (int ` A Un int ` B)" | |
| 133 | "A Int B = nat ` (int ` A Int int ` B)" | |
| 134 |     "{x. P x} = nat ` {x. x >= 0 & P(nat x)}"
 | |
| 135 |     "{..n} = nat ` {0..int n}"
 | |
| 136 |     "{m..n} = nat ` {int m..int n}"  (* need all variants of these! *)
 | |
| 137 | apply (rule card_image [symmetric]) | |
| 138 | apply (auto simp add: inj_on_def image_def) | |
| 139 | apply (rule_tac x = "int x" in bexI) | |
| 140 | apply auto | |
| 141 | apply (rule_tac x = "int x" in bexI) | |
| 142 | apply auto | |
| 143 | apply (rule_tac x = "int x" in bexI) | |
| 144 | apply auto | |
| 145 | apply (rule_tac x = "int x" in exI) | |
| 146 | apply auto | |
| 147 | apply (rule_tac x = "int x" in bexI) | |
| 148 | apply auto | |
| 149 | apply (rule_tac x = "int x" in bexI) | |
| 150 | apply auto | |
| 151 | done | |
| 152 | ||
| 153 | lemma transfer_nat_int_set_function_closures: | |
| 154 |     "nat_set {}"
 | |
| 155 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)" | |
| 156 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)" | |
| 157 |     "x >= 0 \<Longrightarrow> nat_set {x..y}"
 | |
| 158 |     "nat_set {x. x >= 0 & P x}"
 | |
| 159 | "nat_set (int ` C)" | |
| 160 | "nat_set A \<Longrightarrow> x : A \<Longrightarrow> x >= 0" (* does it hurt to turn this on? *) | |
| 161 | unfolding nat_set_def apply auto | |
| 162 | done | |
| 163 | ||
| 164 | lemma transfer_nat_int_set_relations: | |
| 165 | "(finite A) = (finite (int ` A))" | |
| 166 | "(x : A) = (int x : int ` A)" | |
| 167 | "(A = B) = (int ` A = int ` B)" | |
| 168 | "(A < B) = (int ` A < int ` B)" | |
| 169 | "(A <= B) = (int ` A <= int ` B)" | |
| 170 | ||
| 171 | apply (rule iffI) | |
| 172 | apply (erule finite_imageI) | |
| 173 | apply (erule finite_imageD) | |
| 174 | apply (auto simp add: image_def expand_set_eq inj_on_def) | |
| 175 | apply (drule_tac x = "int x" in spec, auto) | |
| 176 | apply (drule_tac x = "int x" in spec, auto) | |
| 177 | apply (drule_tac x = "int x" in spec, auto) | |
| 178 | done | |
| 179 | ||
| 180 | lemma transfer_nat_int_set_return_embed: "nat_set A \<Longrightarrow> | |
| 181 | (int ` nat ` A = A)" | |
| 182 | by (auto simp add: nat_set_def image_def) | |
| 183 | ||
| 184 | lemma transfer_nat_int_set_cong: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 185 |     {(x::int). x >= 0 & P x} = {x. x >= 0 & P' x}"
 | |
| 186 | by auto | |
| 187 | ||
| 188 | declare TransferMorphism_nat_int[transfer add | |
| 189 | return: transfer_nat_int_set_functions | |
| 190 | transfer_nat_int_set_function_closures | |
| 191 | transfer_nat_int_set_relations | |
| 192 | transfer_nat_int_set_return_embed | |
| 193 | cong: transfer_nat_int_set_cong | |
| 194 | ] | |
| 195 | ||
| 196 | ||
| 197 | (* setsum and setprod *) | |
| 198 | ||
| 199 | (* this handles the case where the *domain* of f is nat *) | |
| 200 | lemma transfer_nat_int_sum_prod: | |
| 201 | "setsum f A = setsum (%x. f (nat x)) (int ` A)" | |
| 202 | "setprod f A = setprod (%x. f (nat x)) (int ` A)" | |
| 203 | apply (subst setsum_reindex) | |
| 204 | apply (unfold inj_on_def, auto) | |
| 205 | apply (subst setprod_reindex) | |
| 206 | apply (unfold inj_on_def o_def, auto) | |
| 207 | done | |
| 208 | ||
| 209 | (* this handles the case where the *range* of f is nat *) | |
| 210 | lemma transfer_nat_int_sum_prod2: | |
| 211 | "setsum f A = nat(setsum (%x. int (f x)) A)" | |
| 212 | "setprod f A = nat(setprod (%x. int (f x)) A)" | |
| 213 | apply (subst int_setsum [symmetric]) | |
| 214 | apply auto | |
| 215 | apply (subst int_setprod [symmetric]) | |
| 216 | apply auto | |
| 217 | done | |
| 218 | ||
| 219 | lemma transfer_nat_int_sum_prod_closure: | |
| 220 | "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0" | |
| 221 | "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0" | |
| 222 | unfolding nat_set_def | |
| 223 | apply (rule setsum_nonneg) | |
| 224 | apply auto | |
| 225 | apply (rule setprod_nonneg) | |
| 226 | apply auto | |
| 227 | done | |
| 228 | ||
| 229 | (* this version doesn't work, even with nat_set A \<Longrightarrow> | |
| 230 | x : A \<Longrightarrow> x >= 0 turned on. Why not? | |
| 231 | ||
| 232 | also: what does =simp=> do? | |
| 233 | ||
| 234 | lemma transfer_nat_int_sum_prod_closure: | |
| 235 | "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0" | |
| 236 | "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0" | |
| 237 | unfolding nat_set_def simp_implies_def | |
| 238 | apply (rule setsum_nonneg) | |
| 239 | apply auto | |
| 240 | apply (rule setprod_nonneg) | |
| 241 | apply auto | |
| 242 | done | |
| 243 | *) | |
| 244 | ||
| 245 | (* Making A = B in this lemma doesn't work. Why not? | |
| 246 | Also, why aren't setsum_cong and setprod_cong enough, | |
| 247 | with the previously mentioned rule turned on? *) | |
| 248 | ||
| 249 | lemma transfer_nat_int_sum_prod_cong: | |
| 250 | "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow> | |
| 251 | setsum f A = setsum g B" | |
| 252 | "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow> | |
| 253 | setprod f A = setprod g B" | |
| 254 | unfolding nat_set_def | |
| 255 | apply (subst setsum_cong, assumption) | |
| 256 | apply auto [2] | |
| 257 | apply (subst setprod_cong, assumption, auto) | |
| 258 | done | |
| 259 | ||
| 260 | declare TransferMorphism_nat_int[transfer add | |
| 261 | return: transfer_nat_int_sum_prod transfer_nat_int_sum_prod2 | |
| 262 | transfer_nat_int_sum_prod_closure | |
| 263 | cong: transfer_nat_int_sum_prod_cong] | |
| 264 | ||
| 265 | (* lists *) | |
| 266 | ||
| 267 | definition | |
| 268 | embed_list :: "nat list \<Rightarrow> int list" | |
| 269 | where | |
| 270 | "embed_list l = map int l"; | |
| 271 | ||
| 272 | definition | |
| 273 | nat_list :: "int list \<Rightarrow> bool" | |
| 274 | where | |
| 275 | "nat_list l = nat_set (set l)"; | |
| 276 | ||
| 277 | definition | |
| 278 | return_list :: "int list \<Rightarrow> nat list" | |
| 279 | where | |
| 280 | "return_list l = map nat l"; | |
| 281 | ||
| 282 | thm nat_0_le; | |
| 283 | ||
| 284 | lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow> | |
| 285 | embed_list (return_list l) = l"; | |
| 286 | unfolding embed_list_def return_list_def nat_list_def nat_set_def | |
| 287 | apply (induct l); | |
| 288 | apply auto; | |
| 289 | done; | |
| 290 | ||
| 291 | lemma transfer_nat_int_list_functions: | |
| 292 | "l @ m = return_list (embed_list l @ embed_list m)" | |
| 293 | "[] = return_list []"; | |
| 294 | unfolding return_list_def embed_list_def; | |
| 295 | apply auto; | |
| 296 | apply (induct l, auto); | |
| 297 | apply (induct m, auto); | |
| 298 | done; | |
| 299 | ||
| 300 | (* | |
| 301 | lemma transfer_nat_int_fold1: "fold f l x = | |
| 302 | fold (%x. f (nat x)) (embed_list l) x"; | |
| 303 | *) | |
| 304 | ||
| 305 | ||
| 306 | ||
| 307 | ||
| 308 | subsection {* Set up transfer from int to nat *}
 | |
| 309 | ||
| 310 | (* set up transfer direction *) | |
| 311 | ||
| 312 | lemma TransferMorphism_int_nat: "TransferMorphism int (UNIV :: nat set)" | |
| 313 | by (simp add: TransferMorphism_def) | |
| 314 | ||
| 315 | declare TransferMorphism_int_nat[transfer add | |
| 316 | mode: manual | |
| 317 | (* labels: int-nat *) | |
| 318 | return: nat_int | |
| 319 | ] | |
| 320 | ||
| 321 | ||
| 322 | (* basic functions and relations *) | |
| 323 | ||
| 324 | definition | |
| 325 | is_nat :: "int \<Rightarrow> bool" | |
| 326 | where | |
| 327 | "is_nat x = (x >= 0)" | |
| 328 | ||
| 329 | lemma transfer_int_nat_numerals: | |
| 330 | "0 = int 0" | |
| 331 | "1 = int 1" | |
| 332 | "2 = int 2" | |
| 333 | "3 = int 3" | |
| 334 | by auto | |
| 335 | ||
| 336 | lemma transfer_int_nat_functions: | |
| 337 | "(int x) + (int y) = int (x + y)" | |
| 338 | "(int x) * (int y) = int (x * y)" | |
| 339 | "tsub (int x) (int y) = int (x - y)" | |
| 340 | "(int x)^n = int (x^n)" | |
| 341 | "(int x) div (int y) = int (x div y)" | |
| 342 | "(int x) mod (int y) = int (x mod y)" | |
| 343 | by (auto simp add: int_mult tsub_def int_power zdiv_int zmod_int) | |
| 344 | ||
| 345 | lemma transfer_int_nat_function_closures: | |
| 346 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)" | |
| 347 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)" | |
| 348 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)" | |
| 349 | "is_nat x \<Longrightarrow> is_nat (x^n)" | |
| 350 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)" | |
| 351 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)" | |
| 352 | "is_nat 0" | |
| 353 | "is_nat 1" | |
| 354 | "is_nat 2" | |
| 355 | "is_nat 3" | |
| 356 | "is_nat (int z)" | |
| 357 | by (simp_all only: is_nat_def transfer_nat_int_function_closures) | |
| 358 | ||
| 359 | lemma transfer_int_nat_relations: | |
| 360 | "(int x = int y) = (x = y)" | |
| 361 | "(int x < int y) = (x < y)" | |
| 362 | "(int x <= int y) = (x <= y)" | |
| 363 | "(int x dvd int y) = (x dvd y)" | |
| 364 | "(even (int x)) = (even x)" | |
| 365 | by (auto simp add: zdvd_int even_nat_def) | |
| 366 | ||
| 32121 | 367 | lemma UNIV_apply: | 
| 368 | "UNIV x = True" | |
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
32121diff
changeset | 369 | by (simp add: top_fun_eq top_bool_eq) | 
| 32121 | 370 | |
| 31708 | 371 | declare TransferMorphism_int_nat[transfer add return: | 
| 372 | transfer_int_nat_numerals | |
| 373 | transfer_int_nat_functions | |
| 374 | transfer_int_nat_function_closures | |
| 375 | transfer_int_nat_relations | |
| 32121 | 376 | UNIV_apply | 
| 31708 | 377 | ] | 
| 378 | ||
| 379 | ||
| 380 | (* first-order quantifiers *) | |
| 381 | ||
| 382 | lemma transfer_int_nat_quantifiers: | |
| 383 | "(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))" | |
| 384 | "(EX (x::int) >= 0. P x) = (EX (x::nat). P (int x))" | |
| 385 | apply (subst all_nat) | |
| 386 | apply auto [1] | |
| 387 | apply (subst ex_nat) | |
| 388 | apply auto | |
| 389 | done | |
| 390 | ||
| 391 | declare TransferMorphism_int_nat[transfer add | |
| 392 | return: transfer_int_nat_quantifiers] | |
| 393 | ||
| 394 | ||
| 395 | (* if *) | |
| 396 | ||
| 397 | lemma int_if_cong: "(if P then (int x) else (int y)) = | |
| 398 | int (if P then x else y)" | |
| 399 | by auto | |
| 400 | ||
| 401 | declare TransferMorphism_int_nat [transfer add return: int_if_cong] | |
| 402 | ||
| 403 | ||
| 404 | ||
| 405 | (* operations with sets *) | |
| 406 | ||
| 407 | lemma transfer_int_nat_set_functions: | |
| 408 | "nat_set A \<Longrightarrow> card A = card (nat ` A)" | |
| 409 |     "{} = int ` ({}::nat set)"
 | |
| 410 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)" | |
| 411 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Int B = int ` (nat ` A Int nat ` B)" | |
| 412 |     "{x. x >= 0 & P x} = int ` {x. P(int x)}"
 | |
| 413 |     "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
 | |
| 414 | (* need all variants of these! *) | |
| 415 | by (simp_all only: is_nat_def transfer_nat_int_set_functions | |
| 416 | transfer_nat_int_set_function_closures | |
| 417 | transfer_nat_int_set_return_embed nat_0_le | |
| 418 | cong: transfer_nat_int_set_cong) | |
| 419 | ||
| 420 | lemma transfer_int_nat_set_function_closures: | |
| 421 |     "nat_set {}"
 | |
| 422 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)" | |
| 423 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)" | |
| 424 |     "is_nat x \<Longrightarrow> nat_set {x..y}"
 | |
| 425 |     "nat_set {x. x >= 0 & P x}"
 | |
| 426 | "nat_set (int ` C)" | |
| 427 | "nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x" | |
| 428 | by (simp_all only: transfer_nat_int_set_function_closures is_nat_def) | |
| 429 | ||
| 430 | lemma transfer_int_nat_set_relations: | |
| 431 | "nat_set A \<Longrightarrow> finite A = finite (nat ` A)" | |
| 432 | "is_nat x \<Longrightarrow> nat_set A \<Longrightarrow> (x : A) = (nat x : nat ` A)" | |
| 433 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A = B) = (nat ` A = nat ` B)" | |
| 434 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A < B) = (nat ` A < nat ` B)" | |
| 435 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A <= B) = (nat ` A <= nat ` B)" | |
| 436 | by (simp_all only: is_nat_def transfer_nat_int_set_relations | |
| 437 | transfer_nat_int_set_return_embed nat_0_le) | |
| 438 | ||
| 439 | lemma transfer_int_nat_set_return_embed: "nat ` int ` A = A" | |
| 440 | by (simp only: transfer_nat_int_set_relations | |
| 441 | transfer_nat_int_set_function_closures | |
| 442 | transfer_nat_int_set_return_embed nat_0_le) | |
| 443 | ||
| 444 | lemma transfer_int_nat_set_cong: "(!!x. P x = P' x) \<Longrightarrow> | |
| 445 |     {(x::nat). P x} = {x. P' x}"
 | |
| 446 | by auto | |
| 447 | ||
| 448 | declare TransferMorphism_int_nat[transfer add | |
| 449 | return: transfer_int_nat_set_functions | |
| 450 | transfer_int_nat_set_function_closures | |
| 451 | transfer_int_nat_set_relations | |
| 452 | transfer_int_nat_set_return_embed | |
| 453 | cong: transfer_int_nat_set_cong | |
| 454 | ] | |
| 455 | ||
| 456 | ||
| 457 | (* setsum and setprod *) | |
| 458 | ||
| 459 | (* this handles the case where the *domain* of f is int *) | |
| 460 | lemma transfer_int_nat_sum_prod: | |
| 461 | "nat_set A \<Longrightarrow> setsum f A = setsum (%x. f (int x)) (nat ` A)" | |
| 462 | "nat_set A \<Longrightarrow> setprod f A = setprod (%x. f (int x)) (nat ` A)" | |
| 463 | apply (subst setsum_reindex) | |
| 464 | apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff) | |
| 465 | apply (subst setprod_reindex) | |
| 466 | apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff | |
| 467 | cong: setprod_cong) | |
| 468 | done | |
| 469 | ||
| 470 | (* this handles the case where the *range* of f is int *) | |
| 471 | lemma transfer_int_nat_sum_prod2: | |
| 472 | "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> setsum f A = int(setsum (%x. nat (f x)) A)" | |
| 473 | "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> | |
| 474 | setprod f A = int(setprod (%x. nat (f x)) A)" | |
| 475 | unfolding is_nat_def | |
| 476 | apply (subst int_setsum, auto) | |
| 477 | apply (subst int_setprod, auto simp add: cong: setprod_cong) | |
| 478 | done | |
| 479 | ||
| 480 | declare TransferMorphism_int_nat[transfer add | |
| 481 | return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2 | |
| 482 | cong: setsum_cong setprod_cong] | |
| 483 | ||
| 484 | end |