| author | paulson <lp15@cam.ac.uk> | 
| Fri, 18 Feb 2022 21:40:01 +0000 | |
| changeset 75101 | f0e2023f361a | 
| parent 69913 | ca515cf61651 | 
| child 81586 | 257f93d40d7c | 
| permissions | -rw-r--r-- | 
| 11608 | 1 | (* Title: HOL/Typedef.thy | 
| 2 | Author: Markus Wenzel, TU Munich | |
| 11743 | 3 | *) | 
| 11608 | 4 | |
| 60758 | 5 | section \<open>HOL type definitions\<close> | 
| 11608 | 6 | |
| 15131 | 7 | theory Typedef | 
| 15140 | 8 | imports Set | 
| 63434 | 9 | keywords | 
| 69913 | 10 | "typedef" :: thy_goal_defn and | 
| 63434 | 11 | "morphisms" :: quasi_command | 
| 15131 | 12 | begin | 
| 11608 | 13 | |
| 13412 | 14 | locale type_definition = | 
| 15 | fixes Rep and Abs and A | |
| 16 | assumes Rep: "Rep x \<in> A" | |
| 17 | and Rep_inverse: "Abs (Rep x) = x" | |
| 61102 | 18 | and Abs_inverse: "y \<in> A \<Longrightarrow> Rep (Abs y) = y" | 
| 61799 | 19 | \<comment> \<open>This will be axiomatized for each typedef!\<close> | 
| 23247 | 20 | begin | 
| 11608 | 21 | |
| 61102 | 22 | lemma Rep_inject: "Rep x = Rep y \<longleftrightarrow> x = y" | 
| 13412 | 23 | proof | 
| 24 | assume "Rep x = Rep y" | |
| 23710 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 25 | then have "Abs (Rep x) = Abs (Rep y)" by (simp only:) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 26 | moreover have "Abs (Rep x) = x" by (rule Rep_inverse) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 27 | moreover have "Abs (Rep y) = y" by (rule Rep_inverse) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 28 | ultimately show "x = y" by simp | 
| 13412 | 29 | next | 
| 30 | assume "x = y" | |
| 61102 | 31 | then show "Rep x = Rep y" by (simp only:) | 
| 13412 | 32 | qed | 
| 11608 | 33 | |
| 23247 | 34 | lemma Abs_inject: | 
| 61102 | 35 | assumes "x \<in> A" and "y \<in> A" | 
| 36 | shows "Abs x = Abs y \<longleftrightarrow> x = y" | |
| 13412 | 37 | proof | 
| 38 | assume "Abs x = Abs y" | |
| 23710 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 39 | then have "Rep (Abs x) = Rep (Abs y)" by (simp only:) | 
| 61102 | 40 | moreover from \<open>x \<in> A\<close> have "Rep (Abs x) = x" by (rule Abs_inverse) | 
| 41 | moreover from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse) | |
| 23710 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 42 | ultimately show "x = y" by simp | 
| 13412 | 43 | next | 
| 44 | assume "x = y" | |
| 61102 | 45 | then show "Abs x = Abs y" by (simp only:) | 
| 11608 | 46 | qed | 
| 47 | ||
| 23247 | 48 | lemma Rep_cases [cases set]: | 
| 61102 | 49 | assumes "y \<in> A" | 
| 50 | and hyp: "\<And>x. y = Rep x \<Longrightarrow> P" | |
| 13412 | 51 | shows P | 
| 52 | proof (rule hyp) | |
| 61102 | 53 | from \<open>y \<in> A\<close> have "Rep (Abs y) = y" by (rule Abs_inverse) | 
| 54 | then show "y = Rep (Abs y)" .. | |
| 11608 | 55 | qed | 
| 56 | ||
| 23247 | 57 | lemma Abs_cases [cases type]: | 
| 61102 | 58 | assumes r: "\<And>y. x = Abs y \<Longrightarrow> y \<in> A \<Longrightarrow> P" | 
| 13412 | 59 | shows P | 
| 60 | proof (rule r) | |
| 61 | have "Abs (Rep x) = x" by (rule Rep_inverse) | |
| 61102 | 62 | then show "x = Abs (Rep x)" .. | 
| 13412 | 63 | show "Rep x \<in> A" by (rule Rep) | 
| 11608 | 64 | qed | 
| 65 | ||
| 23247 | 66 | lemma Rep_induct [induct set]: | 
| 13412 | 67 | assumes y: "y \<in> A" | 
| 61102 | 68 | and hyp: "\<And>x. P (Rep x)" | 
| 13412 | 69 | shows "P y" | 
| 11608 | 70 | proof - | 
| 13412 | 71 | have "P (Rep (Abs y))" by (rule hyp) | 
| 23710 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 72 | moreover from y have "Rep (Abs y) = y" by (rule Abs_inverse) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 73 | ultimately show "P y" by simp | 
| 11608 | 74 | qed | 
| 75 | ||
| 23247 | 76 | lemma Abs_induct [induct type]: | 
| 61102 | 77 | assumes r: "\<And>y. y \<in> A \<Longrightarrow> P (Abs y)" | 
| 13412 | 78 | shows "P x" | 
| 11608 | 79 | proof - | 
| 13412 | 80 | have "Rep x \<in> A" by (rule Rep) | 
| 23710 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 81 | then have "P (Abs (Rep x))" by (rule r) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 82 | moreover have "Abs (Rep x) = x" by (rule Rep_inverse) | 
| 
a8ac2305eaf2
removed proof dependency on transitivity theorems
 haftmann parents: 
23433diff
changeset | 83 | ultimately show "P x" by simp | 
| 11608 | 84 | qed | 
| 85 | ||
| 27295 | 86 | lemma Rep_range: "range Rep = A" | 
| 24269 | 87 | proof | 
| 61102 | 88 | show "range Rep \<subseteq> A" using Rep by (auto simp add: image_def) | 
| 89 | show "A \<subseteq> range Rep" | |
| 23433 | 90 | proof | 
| 61102 | 91 | fix x assume "x \<in> A" | 
| 92 | then have "x = Rep (Abs x)" by (rule Abs_inverse [symmetric]) | |
| 93 | then show "x \<in> range Rep" by (rule range_eqI) | |
| 23433 | 94 | qed | 
| 95 | qed | |
| 96 | ||
| 27295 | 97 | lemma Abs_image: "Abs ` A = UNIV" | 
| 98 | proof | |
| 61102 | 99 | show "Abs ` A \<subseteq> UNIV" by (rule subset_UNIV) | 
| 100 | show "UNIV \<subseteq> Abs ` A" | |
| 27295 | 101 | proof | 
| 102 | fix x | |
| 103 | have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) | |
| 61102 | 104 | moreover have "Rep x \<in> A" by (rule Rep) | 
| 105 | ultimately show "x \<in> Abs ` A" by (rule image_eqI) | |
| 27295 | 106 | qed | 
| 107 | qed | |
| 108 | ||
| 23247 | 109 | end | 
| 110 | ||
| 69605 | 111 | ML_file \<open>Tools/typedef.ML\<close> | 
| 11608 | 112 | |
| 113 | end |