| author | haftmann | 
| Sat, 31 Dec 2016 08:12:31 +0100 | |
| changeset 64715 | 33d5fa0ce6e5 | 
| parent 64272 | f76b6dda2e56 | 
| child 66795 | 420d0080545f | 
| permissions | -rw-r--r-- | 
| 32554 | 1 | (* Authors: Jeremy Avigad and Amine Chaieb *) | 
| 31708 | 2 | |
| 60758 | 3 | section \<open>Generic transfer machinery; specific transfer from nats to ints and back.\<close> | 
| 31708 | 4 | |
| 32558 | 5 | theory Nat_Transfer | 
| 47255 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 huffman parents: 
42870diff
changeset | 6 | imports Int | 
| 31708 | 7 | begin | 
| 8 | ||
| 60758 | 9 | subsection \<open>Generic transfer machinery\<close> | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 10 | |
| 35821 | 11 | definition transfer_morphism:: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 42870 
36abaf4cce1f
clarified vacuous nature of predicate "transfer_morphism" -- equivalent to previous definiton
 krauss parents: 
39302diff
changeset | 12 | where "transfer_morphism f A \<longleftrightarrow> True" | 
| 35644 | 13 | |
| 42870 
36abaf4cce1f
clarified vacuous nature of predicate "transfer_morphism" -- equivalent to previous definiton
 krauss parents: 
39302diff
changeset | 14 | lemma transfer_morphismI[intro]: "transfer_morphism f A" | 
| 
36abaf4cce1f
clarified vacuous nature of predicate "transfer_morphism" -- equivalent to previous definiton
 krauss parents: 
39302diff
changeset | 15 | by (simp add: transfer_morphism_def) | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 16 | |
| 48891 | 17 | ML_file "Tools/legacy_transfer.ML" | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 18 | |
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 19 | |
| 60758 | 20 | subsection \<open>Set up transfer from nat to int\<close> | 
| 31708 | 21 | |
| 60758 | 22 | text \<open>set up transfer direction\<close> | 
| 31708 | 23 | |
| 42870 
36abaf4cce1f
clarified vacuous nature of predicate "transfer_morphism" -- equivalent to previous definiton
 krauss parents: 
39302diff
changeset | 24 | lemma transfer_morphism_nat_int: "transfer_morphism nat (op <= (0::int))" .. | 
| 31708 | 25 | |
| 35683 | 26 | declare transfer_morphism_nat_int [transfer add | 
| 27 | mode: manual | |
| 31708 | 28 | return: nat_0_le | 
| 35683 | 29 | labels: nat_int | 
| 31708 | 30 | ] | 
| 31 | ||
| 60758 | 32 | text \<open>basic functions and relations\<close> | 
| 31708 | 33 | |
| 35683 | 34 | lemma transfer_nat_int_numerals [transfer key: transfer_morphism_nat_int]: | 
| 31708 | 35 | "(0::nat) = nat 0" | 
| 36 | "(1::nat) = nat 1" | |
| 37 | "(2::nat) = nat 2" | |
| 38 | "(3::nat) = nat 3" | |
| 39 | by auto | |
| 40 | ||
| 41 | definition | |
| 42 | tsub :: "int \<Rightarrow> int \<Rightarrow> int" | |
| 43 | where | |
| 44 | "tsub x y = (if x >= y then x - y else 0)" | |
| 45 | ||
| 46 | lemma tsub_eq: "x >= y \<Longrightarrow> tsub x y = x - y" | |
| 47 | by (simp add: tsub_def) | |
| 48 | ||
| 35683 | 49 | lemma transfer_nat_int_functions [transfer key: transfer_morphism_nat_int]: | 
| 31708 | 50 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) + (nat y) = nat (x + y)" | 
| 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 (tsub x y)" | |
| 53 | "(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)" | |
| 54 | by (auto simp add: eq_nat_nat_iff nat_mult_distrib | |
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 55 | nat_power_eq tsub_def) | 
| 31708 | 56 | |
| 35683 | 57 | lemma transfer_nat_int_function_closures [transfer key: transfer_morphism_nat_int]: | 
| 31708 | 58 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0" | 
| 59 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0" | |
| 60 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0" | |
| 61 | "(x::int) >= 0 \<Longrightarrow> x^n >= 0" | |
| 62 | "(0::int) >= 0" | |
| 63 | "(1::int) >= 0" | |
| 64 | "(2::int) >= 0" | |
| 65 | "(3::int) >= 0" | |
| 66 | "int z >= 0" | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 67 | by (auto simp add: zero_le_mult_iff tsub_def) | 
| 31708 | 68 | |
| 35683 | 69 | lemma transfer_nat_int_relations [transfer key: transfer_morphism_nat_int]: | 
| 31708 | 70 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | 
| 71 | (nat (x::int) = nat y) = (x = y)" | |
| 72 | "x >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> | |
| 73 | (nat (x::int) < nat y) = (x < y)" | |
| 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) dvd nat y) = (x dvd y)" | |
| 32558 | 78 | by (auto simp add: zdvd_int) | 
| 31708 | 79 | |
| 80 | ||
| 60758 | 81 | text \<open>first-order quantifiers\<close> | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 82 | |
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 83 | lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))" | 
| 63648 | 84 | by (simp split: split_nat) | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 85 | |
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 86 | lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))" | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 87 | proof | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 88 | assume "\<exists>x. P x" | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 89 | then obtain x where "P x" .. | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 90 | then have "int x \<ge> 0 \<and> P (nat (int x))" by simp | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 91 | then show "\<exists>x\<ge>0. P (nat x)" .. | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 92 | next | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 93 | assume "\<exists>x\<ge>0. P (nat x)" | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 94 | then show "\<exists>x. P x" by auto | 
| 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 95 | qed | 
| 31708 | 96 | |
| 35683 | 97 | lemma transfer_nat_int_quantifiers [transfer key: transfer_morphism_nat_int]: | 
| 31708 | 98 | "(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))" | 
| 99 | "(EX (x::nat). P x) = (EX (x::int). x >= 0 & P (nat x))" | |
| 100 | by (rule all_nat, rule ex_nat) | |
| 101 | ||
| 102 | (* should we restrict these? *) | |
| 103 | lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 104 | (ALL x. Q x \<longrightarrow> P x) = (ALL x. Q x \<longrightarrow> P' x)" | |
| 105 | by auto | |
| 106 | ||
| 107 | lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 108 | (EX x. Q x \<and> P x) = (EX x. Q x \<and> P' x)" | |
| 109 | by auto | |
| 110 | ||
| 35644 | 111 | declare transfer_morphism_nat_int [transfer add | 
| 31708 | 112 | cong: all_cong ex_cong] | 
| 113 | ||
| 114 | ||
| 60758 | 115 | text \<open>if\<close> | 
| 31708 | 116 | |
| 35683 | 117 | lemma nat_if_cong [transfer key: transfer_morphism_nat_int]: | 
| 118 | "(if P then (nat x) else (nat y)) = nat (if P then x else y)" | |
| 31708 | 119 | by auto | 
| 120 | ||
| 121 | ||
| 60758 | 122 | text \<open>operations with sets\<close> | 
| 31708 | 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 | apply (rule card_image [symmetric]) | |
| 136 | apply (auto simp add: inj_on_def image_def) | |
| 137 | apply (rule_tac x = "int x" in bexI) | |
| 138 | apply auto | |
| 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 exI) | |
| 144 | apply auto | |
| 145 | done | |
| 146 | ||
| 147 | lemma transfer_nat_int_set_function_closures: | |
| 148 |     "nat_set {}"
 | |
| 149 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)" | |
| 150 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)" | |
| 151 |     "nat_set {x. x >= 0 & P x}"
 | |
| 152 | "nat_set (int ` C)" | |
| 153 | "nat_set A \<Longrightarrow> x : A \<Longrightarrow> x >= 0" (* does it hurt to turn this on? *) | |
| 154 | unfolding nat_set_def apply auto | |
| 155 | done | |
| 156 | ||
| 157 | lemma transfer_nat_int_set_relations: | |
| 158 | "(finite A) = (finite (int ` A))" | |
| 159 | "(x : A) = (int x : int ` A)" | |
| 160 | "(A = B) = (int ` A = int ` B)" | |
| 161 | "(A < B) = (int ` A < int ` B)" | |
| 162 | "(A <= B) = (int ` A <= int ` B)" | |
| 163 | apply (rule iffI) | |
| 164 | apply (erule finite_imageI) | |
| 165 | apply (erule finite_imageD) | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 166 | apply (auto simp add: image_def set_eq_iff inj_on_def) | 
| 31708 | 167 | apply (drule_tac x = "int x" in spec, auto) | 
| 168 | apply (drule_tac x = "int x" in spec, auto) | |
| 169 | apply (drule_tac x = "int x" in spec, auto) | |
| 170 | done | |
| 171 | ||
| 172 | lemma transfer_nat_int_set_return_embed: "nat_set A \<Longrightarrow> | |
| 173 | (int ` nat ` A = A)" | |
| 174 | by (auto simp add: nat_set_def image_def) | |
| 175 | ||
| 176 | lemma transfer_nat_int_set_cong: "(!!x. x >= 0 \<Longrightarrow> P x = P' x) \<Longrightarrow> | |
| 177 |     {(x::int). x >= 0 & P x} = {x. x >= 0 & P' x}"
 | |
| 178 | by auto | |
| 179 | ||
| 35644 | 180 | declare transfer_morphism_nat_int [transfer add | 
| 31708 | 181 | return: transfer_nat_int_set_functions | 
| 182 | transfer_nat_int_set_function_closures | |
| 183 | transfer_nat_int_set_relations | |
| 184 | transfer_nat_int_set_return_embed | |
| 185 | cong: transfer_nat_int_set_cong | |
| 186 | ] | |
| 187 | ||
| 188 | ||
| 64272 | 189 | text \<open>sum and prod\<close> | 
| 31708 | 190 | |
| 191 | (* this handles the case where the *domain* of f is nat *) | |
| 192 | lemma transfer_nat_int_sum_prod: | |
| 64267 | 193 | "sum f A = sum (%x. f (nat x)) (int ` A)" | 
| 64272 | 194 | "prod f A = prod (%x. f (nat x)) (int ` A)" | 
| 64267 | 195 | apply (subst sum.reindex) | 
| 31708 | 196 | apply (unfold inj_on_def, auto) | 
| 64272 | 197 | apply (subst prod.reindex) | 
| 31708 | 198 | apply (unfold inj_on_def o_def, auto) | 
| 199 | done | |
| 200 | ||
| 201 | (* this handles the case where the *range* of f is nat *) | |
| 202 | lemma transfer_nat_int_sum_prod2: | |
| 64267 | 203 | "sum f A = nat(sum (%x. int (f x)) A)" | 
| 64272 | 204 | "prod f A = nat(prod (%x. int (f x)) A)" | 
| 64267 | 205 | apply (simp only: int_sum [symmetric] nat_int) | 
| 64272 | 206 | apply (simp only: int_prod [symmetric] nat_int) | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
60758diff
changeset | 207 | done | 
| 31708 | 208 | |
| 209 | lemma transfer_nat_int_sum_prod_closure: | |
| 64267 | 210 | "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> sum f A >= 0" | 
| 64272 | 211 | "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> prod f A >= 0" | 
| 31708 | 212 | unfolding nat_set_def | 
| 64267 | 213 | apply (rule sum_nonneg) | 
| 31708 | 214 | apply auto | 
| 64272 | 215 | apply (rule prod_nonneg) | 
| 31708 | 216 | apply auto | 
| 217 | done | |
| 218 | ||
| 219 | (* this version doesn't work, even with nat_set A \<Longrightarrow> | |
| 220 | x : A \<Longrightarrow> x >= 0 turned on. Why not? | |
| 221 | ||
| 222 | also: what does =simp=> do? | |
| 223 | ||
| 224 | lemma transfer_nat_int_sum_prod_closure: | |
| 64267 | 225 | "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> sum f A >= 0" | 
| 64272 | 226 | "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> prod f A >= 0" | 
| 31708 | 227 | unfolding nat_set_def simp_implies_def | 
| 64267 | 228 | apply (rule sum_nonneg) | 
| 31708 | 229 | apply auto | 
| 64272 | 230 | apply (rule prod_nonneg) | 
| 31708 | 231 | apply auto | 
| 232 | done | |
| 233 | *) | |
| 234 | ||
| 235 | (* Making A = B in this lemma doesn't work. Why not? | |
| 64272 | 236 | Also, why aren't sum.cong and prod.cong enough, | 
| 31708 | 237 | with the previously mentioned rule turned on? *) | 
| 238 | ||
| 239 | lemma transfer_nat_int_sum_prod_cong: | |
| 240 | "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow> | |
| 64267 | 241 | sum f A = sum g B" | 
| 31708 | 242 | "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow> | 
| 64272 | 243 | prod f A = prod g B" | 
| 31708 | 244 | unfolding nat_set_def | 
| 64267 | 245 | apply (subst sum.cong, assumption) | 
| 31708 | 246 | apply auto [2] | 
| 64272 | 247 | apply (subst prod.cong, assumption, auto) | 
| 31708 | 248 | done | 
| 249 | ||
| 35644 | 250 | declare transfer_morphism_nat_int [transfer add | 
| 31708 | 251 | return: transfer_nat_int_sum_prod transfer_nat_int_sum_prod2 | 
| 252 | transfer_nat_int_sum_prod_closure | |
| 253 | cong: transfer_nat_int_sum_prod_cong] | |
| 254 | ||
| 255 | ||
| 60758 | 256 | subsection \<open>Set up transfer from int to nat\<close> | 
| 31708 | 257 | |
| 60758 | 258 | text \<open>set up transfer direction\<close> | 
| 31708 | 259 | |
| 42870 
36abaf4cce1f
clarified vacuous nature of predicate "transfer_morphism" -- equivalent to previous definiton
 krauss parents: 
39302diff
changeset | 260 | lemma transfer_morphism_int_nat: "transfer_morphism int (\<lambda>n. True)" .. | 
| 31708 | 261 | |
| 35644 | 262 | declare transfer_morphism_int_nat [transfer add | 
| 31708 | 263 | mode: manual | 
| 264 | return: nat_int | |
| 35683 | 265 | labels: int_nat | 
| 31708 | 266 | ] | 
| 267 | ||
| 268 | ||
| 60758 | 269 | text \<open>basic functions and relations\<close> | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 270 | |
| 31708 | 271 | definition | 
| 272 | is_nat :: "int \<Rightarrow> bool" | |
| 273 | where | |
| 274 | "is_nat x = (x >= 0)" | |
| 275 | ||
| 276 | lemma transfer_int_nat_numerals: | |
| 277 | "0 = int 0" | |
| 278 | "1 = int 1" | |
| 279 | "2 = int 2" | |
| 280 | "3 = int 3" | |
| 281 | by auto | |
| 282 | ||
| 283 | lemma transfer_int_nat_functions: | |
| 284 | "(int x) + (int y) = int (x + y)" | |
| 285 | "(int x) * (int y) = int (x * y)" | |
| 286 | "tsub (int x) (int y) = int (x - y)" | |
| 287 | "(int x)^n = int (x^n)" | |
| 62348 | 288 | by (auto simp add: of_nat_mult tsub_def of_nat_power) | 
| 31708 | 289 | |
| 290 | lemma transfer_int_nat_function_closures: | |
| 291 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)" | |
| 292 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)" | |
| 293 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)" | |
| 294 | "is_nat x \<Longrightarrow> is_nat (x^n)" | |
| 295 | "is_nat 0" | |
| 296 | "is_nat 1" | |
| 297 | "is_nat 2" | |
| 298 | "is_nat 3" | |
| 299 | "is_nat (int z)" | |
| 300 | by (simp_all only: is_nat_def transfer_nat_int_function_closures) | |
| 301 | ||
| 302 | lemma transfer_int_nat_relations: | |
| 303 | "(int x = int y) = (x = y)" | |
| 304 | "(int x < int y) = (x < y)" | |
| 305 | "(int x <= int y) = (x <= y)" | |
| 306 | "(int x dvd int y) = (x dvd y)" | |
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
32558diff
changeset | 307 | by (auto simp add: zdvd_int) | 
| 32121 | 308 | |
| 35644 | 309 | declare transfer_morphism_int_nat [transfer add return: | 
| 31708 | 310 | transfer_int_nat_numerals | 
| 311 | transfer_int_nat_functions | |
| 312 | transfer_int_nat_function_closures | |
| 313 | transfer_int_nat_relations | |
| 314 | ] | |
| 315 | ||
| 316 | ||
| 60758 | 317 | text \<open>first-order quantifiers\<close> | 
| 31708 | 318 | |
| 319 | lemma transfer_int_nat_quantifiers: | |
| 320 | "(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))" | |
| 321 | "(EX (x::int) >= 0. P x) = (EX (x::nat). P (int x))" | |
| 322 | apply (subst all_nat) | |
| 323 | apply auto [1] | |
| 324 | apply (subst ex_nat) | |
| 325 | apply auto | |
| 326 | done | |
| 327 | ||
| 35644 | 328 | declare transfer_morphism_int_nat [transfer add | 
| 31708 | 329 | return: transfer_int_nat_quantifiers] | 
| 330 | ||
| 331 | ||
| 60758 | 332 | text \<open>if\<close> | 
| 31708 | 333 | |
| 334 | lemma int_if_cong: "(if P then (int x) else (int y)) = | |
| 335 | int (if P then x else y)" | |
| 336 | by auto | |
| 337 | ||
| 35644 | 338 | declare transfer_morphism_int_nat [transfer add return: int_if_cong] | 
| 31708 | 339 | |
| 340 | ||
| 341 | ||
| 60758 | 342 | text \<open>operations with sets\<close> | 
| 31708 | 343 | |
| 344 | lemma transfer_int_nat_set_functions: | |
| 345 | "nat_set A \<Longrightarrow> card A = card (nat ` A)" | |
| 346 |     "{} = int ` ({}::nat set)"
 | |
| 347 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)" | |
| 348 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Int B = int ` (nat ` A Int nat ` B)" | |
| 349 |     "{x. x >= 0 & P x} = int ` {x. P(int x)}"
 | |
| 350 | (* need all variants of these! *) | |
| 351 | by (simp_all only: is_nat_def transfer_nat_int_set_functions | |
| 352 | transfer_nat_int_set_function_closures | |
| 353 | transfer_nat_int_set_return_embed nat_0_le | |
| 354 | cong: transfer_nat_int_set_cong) | |
| 355 | ||
| 356 | lemma transfer_int_nat_set_function_closures: | |
| 357 |     "nat_set {}"
 | |
| 358 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)" | |
| 359 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)" | |
| 360 |     "nat_set {x. x >= 0 & P x}"
 | |
| 361 | "nat_set (int ` C)" | |
| 362 | "nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x" | |
| 363 | by (simp_all only: transfer_nat_int_set_function_closures is_nat_def) | |
| 364 | ||
| 365 | lemma transfer_int_nat_set_relations: | |
| 366 | "nat_set A \<Longrightarrow> finite A = finite (nat ` A)" | |
| 367 | "is_nat x \<Longrightarrow> nat_set A \<Longrightarrow> (x : A) = (nat x : nat ` A)" | |
| 368 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A = B) = (nat ` A = nat ` B)" | |
| 369 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A < B) = (nat ` A < nat ` B)" | |
| 370 | "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> (A <= B) = (nat ` A <= nat ` B)" | |
| 371 | by (simp_all only: is_nat_def transfer_nat_int_set_relations | |
| 372 | transfer_nat_int_set_return_embed nat_0_le) | |
| 373 | ||
| 374 | lemma transfer_int_nat_set_return_embed: "nat ` int ` A = A" | |
| 375 | by (simp only: transfer_nat_int_set_relations | |
| 376 | transfer_nat_int_set_function_closures | |
| 377 | transfer_nat_int_set_return_embed nat_0_le) | |
| 378 | ||
| 379 | lemma transfer_int_nat_set_cong: "(!!x. P x = P' x) \<Longrightarrow> | |
| 380 |     {(x::nat). P x} = {x. P' x}"
 | |
| 381 | by auto | |
| 382 | ||
| 35644 | 383 | declare transfer_morphism_int_nat [transfer add | 
| 31708 | 384 | return: transfer_int_nat_set_functions | 
| 385 | transfer_int_nat_set_function_closures | |
| 386 | transfer_int_nat_set_relations | |
| 387 | transfer_int_nat_set_return_embed | |
| 388 | cong: transfer_int_nat_set_cong | |
| 389 | ] | |
| 390 | ||
| 391 | ||
| 64272 | 392 | text \<open>sum and prod\<close> | 
| 31708 | 393 | |
| 394 | (* this handles the case where the *domain* of f is int *) | |
| 395 | lemma transfer_int_nat_sum_prod: | |
| 64267 | 396 | "nat_set A \<Longrightarrow> sum f A = sum (%x. f (int x)) (nat ` A)" | 
| 64272 | 397 | "nat_set A \<Longrightarrow> prod f A = prod (%x. f (int x)) (nat ` A)" | 
| 64267 | 398 | apply (subst sum.reindex) | 
| 31708 | 399 | apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff) | 
| 64272 | 400 | apply (subst prod.reindex) | 
| 31708 | 401 | apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff | 
| 64272 | 402 | cong: prod.cong) | 
| 31708 | 403 | done | 
| 404 | ||
| 405 | (* this handles the case where the *range* of f is int *) | |
| 406 | lemma transfer_int_nat_sum_prod2: | |
| 64267 | 407 | "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> sum f A = int(sum (%x. nat (f x)) A)" | 
| 31708 | 408 | "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> | 
| 64272 | 409 | prod f A = int(prod (%x. nat (f x)) A)" | 
| 31708 | 410 | unfolding is_nat_def | 
| 64267 | 411 | by (subst int_sum) auto | 
| 31708 | 412 | |
| 35644 | 413 | declare transfer_morphism_int_nat [transfer add | 
| 31708 | 414 | return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2 | 
| 64272 | 415 | cong: sum.cong prod.cong] | 
| 31708 | 416 | |
| 417 | end |