| author | blanchet | 
| Thu, 19 Sep 2013 03:13:13 +0200 | |
| changeset 53730 | f2f6874893df | 
| parent 49834 | b27bbb021df1 | 
| child 58881 | b9556a055632 | 
| permissions | -rw-r--r-- | 
| 35100 
53754ec7360b
renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid clash with new theory Quotient in Main HOL;
 wenzelm parents: 
30738diff
changeset | 1 | (* Title: HOL/Library/Quotient_Type.thy | 
| 10483 | 2 | Author: Markus Wenzel, TU Muenchen | 
| 10250 | 3 | *) | 
| 4 | ||
| 14706 | 5 | header {* Quotient types *}
 | 
| 10250 | 6 | |
| 35100 
53754ec7360b
renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid clash with new theory Quotient in Main HOL;
 wenzelm parents: 
30738diff
changeset | 7 | theory Quotient_Type | 
| 30738 | 8 | imports Main | 
| 15131 | 9 | begin | 
| 10250 | 10 | |
| 11 | text {*
 | |
| 10285 | 12 | We introduce the notion of quotient types over equivalence relations | 
| 22390 | 13 | via type classes. | 
| 10250 | 14 | *} | 
| 15 | ||
| 10285 | 16 | subsection {* Equivalence relations and quotient types *}
 | 
| 10250 | 17 | |
| 18 | text {*
 | |
| 10390 | 19 |  \medskip Type class @{text equiv} models equivalence relations @{text
 | 
| 20 | "\<sim> :: 'a => 'a => bool"}. | |
| 10250 | 21 | *} | 
| 22 | ||
| 29608 | 23 | class eqv = | 
| 25062 | 24 | fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sim>" 50) | 
| 10250 | 25 | |
| 22390 | 26 | class equiv = eqv + | 
| 25062 | 27 | assumes equiv_refl [intro]: "x \<sim> x" | 
| 28 | assumes equiv_trans [trans]: "x \<sim> y \<Longrightarrow> y \<sim> z \<Longrightarrow> x \<sim> z" | |
| 29 | assumes equiv_sym [sym]: "x \<sim> y \<Longrightarrow> y \<sim> x" | |
| 10250 | 30 | |
| 12371 | 31 | lemma equiv_not_sym [sym]: "\<not> (x \<sim> y) ==> \<not> (y \<sim> (x::'a::equiv))" | 
| 10477 | 32 | proof - | 
| 23373 | 33 | assume "\<not> (x \<sim> y)" then show "\<not> (y \<sim> x)" | 
| 10477 | 34 | by (rule contrapos_nn) (rule equiv_sym) | 
| 35 | qed | |
| 36 | ||
| 37 | lemma not_equiv_trans1 [trans]: "\<not> (x \<sim> y) ==> y \<sim> z ==> \<not> (x \<sim> (z::'a::equiv))" | |
| 38 | proof - | |
| 23373 | 39 | assume "\<not> (x \<sim> y)" and "y \<sim> z" | 
| 10477 | 40 | show "\<not> (x \<sim> z)" | 
| 41 | proof | |
| 42 | assume "x \<sim> z" | |
| 23373 | 43 | also from `y \<sim> z` have "z \<sim> y" .. | 
| 10477 | 44 | finally have "x \<sim> y" . | 
| 23373 | 45 | with `\<not> (x \<sim> y)` show False by contradiction | 
| 10477 | 46 | qed | 
| 47 | qed | |
| 48 | ||
| 49 | lemma not_equiv_trans2 [trans]: "x \<sim> y ==> \<not> (y \<sim> z) ==> \<not> (x \<sim> (z::'a::equiv))" | |
| 50 | proof - | |
| 23373 | 51 | assume "\<not> (y \<sim> z)" then have "\<not> (z \<sim> y)" .. | 
| 52 | also assume "x \<sim> y" then have "y \<sim> x" .. | |
| 53 | finally have "\<not> (z \<sim> x)" . then show "(\<not> x \<sim> z)" .. | |
| 10477 | 54 | qed | 
| 55 | ||
| 10250 | 56 | text {*
 | 
| 10285 | 57 |  \medskip The quotient type @{text "'a quot"} consists of all
 | 
| 58 |  \emph{equivalence classes} over elements of the base type @{typ 'a}.
 | |
| 10250 | 59 | *} | 
| 60 | ||
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
35100diff
changeset | 61 | definition "quot = {{x. a \<sim> x} | a::'a::eqv. True}"
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
35100diff
changeset | 62 | |
| 49834 | 63 | typedef 'a quot = "quot :: 'a::eqv set set" | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
35100diff
changeset | 64 | unfolding quot_def by blast | 
| 10250 | 65 | |
| 66 | lemma quotI [intro]: "{x. a \<sim> x} \<in> quot"
 | |
| 18730 | 67 | unfolding quot_def by blast | 
| 10250 | 68 | |
| 69 | lemma quotE [elim]: "R \<in> quot ==> (!!a. R = {x. a \<sim> x} ==> C) ==> C"
 | |
| 18730 | 70 | unfolding quot_def by blast | 
| 10250 | 71 | |
| 72 | text {*
 | |
| 73 | \medskip Abstracted equivalence classes are the canonical | |
| 74 | representation of elements of a quotient type. | |
| 75 | *} | |
| 76 | ||
| 19086 | 77 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19086diff
changeset | 78 |   "class" :: "'a::equiv => 'a quot"  ("\<lfloor>_\<rfloor>") where
 | 
| 19086 | 79 |   "\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}"
 | 
| 10250 | 80 | |
| 10311 | 81 | theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>" | 
| 10278 | 82 | proof (cases A) | 
| 83 | fix R assume R: "A = Abs_quot R" | |
| 23373 | 84 |   assume "R \<in> quot" then have "\<exists>a. R = {x. a \<sim> x}" by blast
 | 
| 10278 | 85 |   with R have "\<exists>a. A = Abs_quot {x. a \<sim> x}" by blast
 | 
| 23373 | 86 | then show ?thesis unfolding class_def . | 
| 10250 | 87 | qed | 
| 88 | ||
| 10311 | 89 | lemma quot_cases [cases type: quot]: "(!!a. A = \<lfloor>a\<rfloor> ==> C) ==> C" | 
| 18730 | 90 | using quot_exhaust by blast | 
| 10250 | 91 | |
| 92 | ||
| 10285 | 93 | subsection {* Equality on quotients *}
 | 
| 10250 | 94 | |
| 95 | text {*
 | |
| 10286 | 96 | Equality of canonical quotient elements coincides with the original | 
| 97 | relation. | |
| 10250 | 98 | *} | 
| 99 | ||
| 12371 | 100 | theorem quot_equality [iff?]: "(\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> b)" | 
| 10285 | 101 | proof | 
| 102 | assume eq: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>" | |
| 103 | show "a \<sim> b" | |
| 104 | proof - | |
| 105 |     from eq have "{x. a \<sim> x} = {x. b \<sim> x}"
 | |
| 10551 | 106 | by (simp only: class_def Abs_quot_inject quotI) | 
| 10285 | 107 | moreover have "a \<sim> a" .. | 
| 108 |     ultimately have "a \<in> {x. b \<sim> x}" by blast
 | |
| 23373 | 109 | then have "b \<sim> a" by blast | 
| 110 | then show ?thesis .. | |
| 10285 | 111 | qed | 
| 112 | next | |
| 10250 | 113 | assume ab: "a \<sim> b" | 
| 10285 | 114 | show "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>" | 
| 115 | proof - | |
| 116 |     have "{x. a \<sim> x} = {x. b \<sim> x}"
 | |
| 117 | proof (rule Collect_cong) | |
| 118 | fix x show "(a \<sim> x) = (b \<sim> x)" | |
| 119 | proof | |
| 120 | from ab have "b \<sim> a" .. | |
| 121 | also assume "a \<sim> x" | |
| 122 | finally show "b \<sim> x" . | |
| 123 | next | |
| 124 | note ab | |
| 125 | also assume "b \<sim> x" | |
| 126 | finally show "a \<sim> x" . | |
| 127 | qed | |
| 10250 | 128 | qed | 
| 23373 | 129 | then show ?thesis by (simp only: class_def) | 
| 10250 | 130 | qed | 
| 131 | qed | |
| 132 | ||
| 133 | ||
| 10285 | 134 | subsection {* Picking representing elements *}
 | 
| 10250 | 135 | |
| 19086 | 136 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19086diff
changeset | 137 | pick :: "'a::equiv quot => 'a" where | 
| 19086 | 138 | "pick A = (SOME a. A = \<lfloor>a\<rfloor>)" | 
| 10250 | 139 | |
| 10285 | 140 | theorem pick_equiv [intro]: "pick \<lfloor>a\<rfloor> \<sim> a" | 
| 10250 | 141 | proof (unfold pick_def) | 
| 142 | show "(SOME x. \<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>) \<sim> a" | |
| 143 | proof (rule someI2) | |
| 144 | show "\<lfloor>a\<rfloor> = \<lfloor>a\<rfloor>" .. | |
| 145 | fix x assume "\<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>" | |
| 23373 | 146 | then have "a \<sim> x" .. then show "x \<sim> a" .. | 
| 10250 | 147 | qed | 
| 148 | qed | |
| 149 | ||
| 10483 | 150 | theorem pick_inverse [intro]: "\<lfloor>pick A\<rfloor> = A" | 
| 10250 | 151 | proof (cases A) | 
| 152 | fix a assume a: "A = \<lfloor>a\<rfloor>" | |
| 23373 | 153 | then have "pick A \<sim> a" by (simp only: pick_equiv) | 
| 154 | then have "\<lfloor>pick A\<rfloor> = \<lfloor>a\<rfloor>" .. | |
| 10250 | 155 | with a show ?thesis by simp | 
| 156 | qed | |
| 157 | ||
| 10285 | 158 | text {*
 | 
| 159 | \medskip The following rules support canonical function definitions | |
| 10483 | 160 | on quotient types (with up to two arguments). Note that the | 
| 161 | stripped-down version without additional conditions is sufficient | |
| 162 | most of the time. | |
| 10285 | 163 | *} | 
| 164 | ||
| 10483 | 165 | theorem quot_cond_function: | 
| 18372 | 166 | assumes eq: "!!X Y. P X Y ==> f X Y == g (pick X) (pick Y)" | 
| 167 | and cong: "!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor> | |
| 168 | ==> P \<lfloor>x\<rfloor> \<lfloor>y\<rfloor> ==> P \<lfloor>x'\<rfloor> \<lfloor>y'\<rfloor> ==> g x y = g x' y'" | |
| 169 | and P: "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>" | |
| 170 | shows "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" | |
| 10473 
4f15b844fea6
separate rules for function/operation definitions;
 wenzelm parents: 
10459diff
changeset | 171 | proof - | 
| 18372 | 172 | from eq and P have "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g (pick \<lfloor>a\<rfloor>) (pick \<lfloor>b\<rfloor>)" by (simp only:) | 
| 10505 | 173 | also have "... = g a b" | 
| 10491 
e4a408728012
quot_cond_function: simplified, support conditional definition;
 wenzelm parents: 
10483diff
changeset | 174 | proof (rule cong) | 
| 10483 | 175 | show "\<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> = \<lfloor>a\<rfloor>" .. | 
| 176 | moreover | |
| 177 | show "\<lfloor>pick \<lfloor>b\<rfloor>\<rfloor> = \<lfloor>b\<rfloor>" .. | |
| 10491 
e4a408728012
quot_cond_function: simplified, support conditional definition;
 wenzelm parents: 
10483diff
changeset | 178 | moreover | 
| 23373 | 179 | show "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>" by (rule P) | 
| 10491 
e4a408728012
quot_cond_function: simplified, support conditional definition;
 wenzelm parents: 
10483diff
changeset | 180 | ultimately show "P \<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> \<lfloor>pick \<lfloor>b\<rfloor>\<rfloor>" by (simp only:) | 
| 10285 | 181 | qed | 
| 182 | finally show ?thesis . | |
| 183 | qed | |
| 184 | ||
| 10483 | 185 | theorem quot_function: | 
| 18372 | 186 | assumes "!!X Y. f X Y == g (pick X) (pick Y)" | 
| 187 | and "!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor> ==> g x y = g x' y'" | |
| 188 | shows "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" | |
| 23394 | 189 | using assms and TrueI | 
| 18372 | 190 | by (rule quot_cond_function) | 
| 10285 | 191 | |
| 10499 | 192 | theorem quot_function': | 
| 193 | "(!!X Y. f X Y == g (pick X) (pick Y)) ==> | |
| 194 | (!!x x' y y'. x \<sim> x' ==> y \<sim> y' ==> g x y = g x' y') ==> | |
| 195 | f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" | |
| 18372 | 196 | by (rule quot_function) (simp_all only: quot_equality) | 
| 10499 | 197 | |
| 10250 | 198 | end |