author | wenzelm |
Tue, 22 Jul 2025 12:02:53 +0200 | |
changeset 82894 | a8e47bd31965 |
parent 81125 | ec121999a9cb |
permissions | -rw-r--r-- |
35762 | 1 |
(* Title: ZF/QPair.thy |
1478 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 3 |
Copyright 1993 University of Cambridge |
4 |
||
13285 | 5 |
Many proofs are borrowed from pair.thy and sum.thy |
6 |
||
7 |
Do we EVER have rank(a) < rank(<a;b>) ? Perhaps if the latter rank |
|
46953 | 8 |
is not a limit ordinal? |
0 | 9 |
*) |
10 |
||
60770 | 11 |
section\<open>Quine-Inspired Ordered Pairs and Disjoint Sums\<close> |
13356 | 12 |
|
16417 | 13 |
theory QPair imports Sum func begin |
13285 | 14 |
|
60770 | 15 |
text\<open>For non-well-founded data |
13356 | 16 |
structures in ZF. Does not precisely follow Quine's construction. Thanks |
17 |
to Thomas Forster for suggesting this approach! |
|
18 |
||
19 |
W. V. Quine, On Ordered Pairs and Relations, in Selected Logic Papers, |
|
20 |
1966. |
|
60770 | 21 |
\<close> |
13356 | 22 |
|
24893 | 23 |
definition |
81125 | 24 |
QPair :: "[i, i] \<Rightarrow> i" (\<open>(\<open>indent=1 notation=\<open>mixfix Quine pair\<close>\<close><_;/ _>)\<close>) |
25 |
where "<a;b> \<equiv> a+b" |
|
3923 | 26 |
|
24893 | 27 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
28 |
qfst :: "i \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
29 |
"qfst(p) \<equiv> THE a. \<exists>b. p=<a;b>" |
13285 | 30 |
|
24893 | 31 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
32 |
qsnd :: "i \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
33 |
"qsnd(p) \<equiv> THE b. \<exists>a. p=<a;b>" |
13285 | 34 |
|
24893 | 35 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
36 |
qsplit :: "[[i, i] \<Rightarrow> 'a, i] \<Rightarrow> 'a::{}" (*for pattern-matching*) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
37 |
"qsplit(c,p) \<equiv> c(qfst(p), qsnd(p))" |
0 | 38 |
|
24893 | 39 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
40 |
qconverse :: "i \<Rightarrow> i" where |
76214 | 41 |
"qconverse(r) \<equiv> {z. w \<in> r, \<exists>x y. w=<x;y> \<and> z=<y;x>}" |
13285 | 42 |
|
24893 | 43 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
44 |
QSigma :: "[i, i \<Rightarrow> i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
45 |
"QSigma(A,B) \<equiv> \<Union>x\<in>A. \<Union>y\<in>B(x). {<x;y>}" |
0 | 46 |
|
929
137035296ad6
Moved declarations of @QSUM and <*> to a syntax section.
lcp
parents:
753
diff
changeset
|
47 |
syntax |
80917 | 48 |
"_QSUM" :: "[idt, i, i] \<Rightarrow> i" (\<open>(\<open>indent=3 notation=\<open>binder QSUM\<in>\<close>\<close>QSUM _ \<in> _./ _)\<close> 10) |
80761 | 49 |
syntax_consts |
50 |
"_QSUM" \<rightleftharpoons> QSigma |
|
0 | 51 |
translations |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
52 |
"QSUM x \<in> A. B" => "CONST QSigma(A, \<lambda>x. B)" |
22808 | 53 |
|
54 |
abbreviation |
|
69587 | 55 |
qprod (infixr \<open><*>\<close> 80) where |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
56 |
"A <*> B \<equiv> QSigma(A, \<lambda>_. B)" |
0 | 57 |
|
24893 | 58 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
59 |
qsum :: "[i,i]\<Rightarrow>i" (infixr \<open><+>\<close> 65) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
60 |
"A <+> B \<equiv> ({0} <*> A) \<union> ({1} <*> B)" |
3923 | 61 |
|
24893 | 62 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
63 |
QInl :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
64 |
"QInl(a) \<equiv> <0;a>" |
13285 | 65 |
|
24893 | 66 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
67 |
QInr :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
68 |
"QInr(b) \<equiv> <1;b>" |
13285 | 69 |
|
24893 | 70 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
71 |
qcase :: "[i\<Rightarrow>i, i\<Rightarrow>i, i]\<Rightarrow>i" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
72 |
"qcase(c,d) \<equiv> qsplit(\<lambda>y z. cond(y, d(z), c(z)))" |
13285 | 73 |
|
74 |
||
60770 | 75 |
subsection\<open>Quine ordered pairing\<close> |
13285 | 76 |
|
77 |
(** Lemmas for showing that <a;b> uniquely determines a and b **) |
|
78 |
||
79 |
lemma QPair_empty [simp]: "<0;0> = 0" |
|
80 |
by (simp add: QPair_def) |
|
81 |
||
76214 | 82 |
lemma QPair_iff [simp]: "<a;b> = <c;d> \<longleftrightarrow> a=c \<and> b=d" |
13285 | 83 |
apply (simp add: QPair_def) |
84 |
apply (rule sum_equal_iff) |
|
85 |
done |
|
86 |
||
45602 | 87 |
lemmas QPair_inject = QPair_iff [THEN iffD1, THEN conjE, elim!] |
13285 | 88 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
89 |
lemma QPair_inject1: "<a;b> = <c;d> \<Longrightarrow> a=c" |
13285 | 90 |
by blast |
91 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
92 |
lemma QPair_inject2: "<a;b> = <c;d> \<Longrightarrow> b=d" |
13285 | 93 |
by blast |
94 |
||
95 |
||
60770 | 96 |
subsubsection\<open>QSigma: Disjoint union of a family of sets |
97 |
Generalizes Cartesian product\<close> |
|
13285 | 98 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
99 |
lemma QSigmaI [intro!]: "\<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> <a;b> \<in> QSigma(A,B)" |
13285 | 100 |
by (simp add: QSigma_def) |
101 |
||
102 |
||
103 |
(** Elimination rules for <a;b>:A*B -- introducing no eigenvariables **) |
|
104 |
||
105 |
lemma QSigmaE [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
106 |
"\<lbrakk>c \<in> QSigma(A,B); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
107 |
\<And>x y.\<lbrakk>x \<in> A; y \<in> B(x); c=<x;y>\<rbrakk> \<Longrightarrow> P |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
108 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 109 |
by (simp add: QSigma_def, blast) |
13285 | 110 |
|
111 |
lemma QSigmaE2 [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
112 |
"\<lbrakk><a;b>: QSigma(A,B); \<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
46953 | 113 |
by (simp add: QSigma_def) |
13285 | 114 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
115 |
lemma QSigmaD1: "<a;b> \<in> QSigma(A,B) \<Longrightarrow> a \<in> A" |
13285 | 116 |
by blast |
117 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
118 |
lemma QSigmaD2: "<a;b> \<in> QSigma(A,B) \<Longrightarrow> b \<in> B(a)" |
13285 | 119 |
by blast |
120 |
||
121 |
lemma QSigma_cong: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
122 |
"\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow> |
13285 | 123 |
QSigma(A,B) = QSigma(A',B')" |
46953 | 124 |
by (simp add: QSigma_def) |
13285 | 125 |
|
126 |
lemma QSigma_empty1 [simp]: "QSigma(0,B) = 0" |
|
127 |
by blast |
|
128 |
||
129 |
lemma QSigma_empty2 [simp]: "A <*> 0 = 0" |
|
130 |
by blast |
|
131 |
||
132 |
||
60770 | 133 |
subsubsection\<open>Projections: qfst, qsnd\<close> |
13285 | 134 |
|
135 |
lemma qfst_conv [simp]: "qfst(<a;b>) = a" |
|
13544 | 136 |
by (simp add: qfst_def) |
13285 | 137 |
|
138 |
lemma qsnd_conv [simp]: "qsnd(<a;b>) = b" |
|
13544 | 139 |
by (simp add: qsnd_def) |
13285 | 140 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
141 |
lemma qfst_type [TC]: "p \<in> QSigma(A,B) \<Longrightarrow> qfst(p) \<in> A" |
13285 | 142 |
by auto |
143 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
144 |
lemma qsnd_type [TC]: "p \<in> QSigma(A,B) \<Longrightarrow> qsnd(p) \<in> B(qfst(p))" |
13285 | 145 |
by auto |
146 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
147 |
lemma QPair_qfst_qsnd_eq: "a \<in> QSigma(A,B) \<Longrightarrow> <qfst(a); qsnd(a)> = a" |
13285 | 148 |
by auto |
149 |
||
150 |
||
60770 | 151 |
subsubsection\<open>Eliminator: qsplit\<close> |
13285 | 152 |
|
153 |
(*A META-equality, so that it applies to higher types as well...*) |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
154 |
lemma qsplit [simp]: "qsplit(\<lambda>x y. c(x,y), <a;b>) \<equiv> c(a,b)" |
13285 | 155 |
by (simp add: qsplit_def) |
156 |
||
157 |
||
158 |
lemma qsplit_type [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
159 |
"\<lbrakk>p \<in> QSigma(A,B); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
160 |
\<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(<x;y>) |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
161 |
\<rbrakk> \<Longrightarrow> qsplit(\<lambda>x y. c(x,y), p) \<in> C(p)" |
46953 | 162 |
by auto |
13285 | 163 |
|
46953 | 164 |
lemma expand_qsplit: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
165 |
"u \<in> A<*>B \<Longrightarrow> R(qsplit(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x;y> \<longrightarrow> R(c(x,y)))" |
13285 | 166 |
apply (simp add: qsplit_def, auto) |
167 |
done |
|
168 |
||
169 |
||
60770 | 170 |
subsubsection\<open>qsplit for predicates: result type o\<close> |
13285 | 171 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
172 |
lemma qsplitI: "R(a,b) \<Longrightarrow> qsplit(R, <a;b>)" |
13285 | 173 |
by (simp add: qsplit_def) |
174 |
||
175 |
||
176 |
lemma qsplitE: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
177 |
"\<lbrakk>qsplit(R,z); z \<in> QSigma(A,B); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
178 |
\<And>x y. \<lbrakk>z = <x;y>; R(x,y)\<rbrakk> \<Longrightarrow> P |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
179 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 180 |
by (simp add: qsplit_def, auto) |
13285 | 181 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
182 |
lemma qsplitD: "qsplit(R,<a;b>) \<Longrightarrow> R(a,b)" |
13285 | 183 |
by (simp add: qsplit_def) |
184 |
||
185 |
||
60770 | 186 |
subsubsection\<open>qconverse\<close> |
13285 | 187 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
188 |
lemma qconverseI [intro!]: "<a;b>:r \<Longrightarrow> <b;a>:qconverse(r)" |
13285 | 189 |
by (simp add: qconverse_def, blast) |
190 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
191 |
lemma qconverseD [elim!]: "<a;b> \<in> qconverse(r) \<Longrightarrow> <b;a> \<in> r" |
13285 | 192 |
by (simp add: qconverse_def, blast) |
193 |
||
194 |
lemma qconverseE [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
195 |
"\<lbrakk>yx \<in> qconverse(r); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
196 |
\<And>x y. \<lbrakk>yx=<y;x>; <x;y>:r\<rbrakk> \<Longrightarrow> P |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
197 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 198 |
by (simp add: qconverse_def, blast) |
13285 | 199 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
200 |
lemma qconverse_qconverse: "r<=QSigma(A,B) \<Longrightarrow> qconverse(qconverse(r)) = r" |
13285 | 201 |
by blast |
202 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
203 |
lemma qconverse_type: "r \<subseteq> A <*> B \<Longrightarrow> qconverse(r) \<subseteq> B <*> A" |
13285 | 204 |
by blast |
205 |
||
206 |
lemma qconverse_prod: "qconverse(A <*> B) = B <*> A" |
|
207 |
by blast |
|
208 |
||
209 |
lemma qconverse_empty: "qconverse(0) = 0" |
|
210 |
by blast |
|
211 |
||
212 |
||
60770 | 213 |
subsection\<open>The Quine-inspired notion of disjoint sum\<close> |
13285 | 214 |
|
215 |
lemmas qsum_defs = qsum_def QInl_def QInr_def qcase_def |
|
216 |
||
217 |
(** Introduction rules for the injections **) |
|
218 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
219 |
lemma QInlI [intro!]: "a \<in> A \<Longrightarrow> QInl(a) \<in> A <+> B" |
13285 | 220 |
by (simp add: qsum_defs, blast) |
1097
01379c46ad2d
Changed definitions so that qsplit is now defined in terms of
lcp
parents:
929
diff
changeset
|
221 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
222 |
lemma QInrI [intro!]: "b \<in> B \<Longrightarrow> QInr(b) \<in> A <+> B" |
13285 | 223 |
by (simp add: qsum_defs, blast) |
224 |
||
225 |
(** Elimination rules **) |
|
226 |
||
227 |
lemma qsumE [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
228 |
"\<lbrakk>u \<in> A <+> B; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
229 |
\<And>x. \<lbrakk>x \<in> A; u=QInl(x)\<rbrakk> \<Longrightarrow> P; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
230 |
\<And>y. \<lbrakk>y \<in> B; u=QInr(y)\<rbrakk> \<Longrightarrow> P |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
231 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 232 |
by (simp add: qsum_defs, blast) |
13285 | 233 |
|
234 |
||
235 |
(** Injection and freeness equivalences, for rewriting **) |
|
236 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
237 |
lemma QInl_iff [iff]: "QInl(a)=QInl(b) \<longleftrightarrow> a=b" |
13285 | 238 |
by (simp add: qsum_defs ) |
239 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
240 |
lemma QInr_iff [iff]: "QInr(a)=QInr(b) \<longleftrightarrow> a=b" |
13285 | 241 |
by (simp add: qsum_defs ) |
242 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
243 |
lemma QInl_QInr_iff [simp]: "QInl(a)=QInr(b) \<longleftrightarrow> False" |
13285 | 244 |
by (simp add: qsum_defs ) |
245 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
246 |
lemma QInr_QInl_iff [simp]: "QInr(b)=QInl(a) \<longleftrightarrow> False" |
13285 | 247 |
by (simp add: qsum_defs ) |
248 |
||
249 |
lemma qsum_empty [simp]: "0<+>0 = 0" |
|
250 |
by (simp add: qsum_defs ) |
|
251 |
||
252 |
(*Injection and freeness rules*) |
|
253 |
||
45602 | 254 |
lemmas QInl_inject = QInl_iff [THEN iffD1] |
255 |
lemmas QInr_inject = QInr_iff [THEN iffD1] |
|
13823 | 256 |
lemmas QInl_neq_QInr = QInl_QInr_iff [THEN iffD1, THEN FalseE, elim!] |
257 |
lemmas QInr_neq_QInl = QInr_QInl_iff [THEN iffD1, THEN FalseE, elim!] |
|
13285 | 258 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
259 |
lemma QInlD: "QInl(a): A<+>B \<Longrightarrow> a \<in> A" |
13285 | 260 |
by blast |
261 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
262 |
lemma QInrD: "QInr(b): A<+>B \<Longrightarrow> b \<in> B" |
13285 | 263 |
by blast |
264 |
||
265 |
(** <+> is itself injective... who cares?? **) |
|
266 |
||
267 |
lemma qsum_iff: |
|
76214 | 268 |
"u \<in> A <+> B \<longleftrightarrow> (\<exists>x. x \<in> A \<and> u=QInl(x)) | (\<exists>y. y \<in> B \<and> u=QInr(y))" |
13356 | 269 |
by blast |
13285 | 270 |
|
76214 | 271 |
lemma qsum_subset_iff: "A <+> B \<subseteq> C <+> D \<longleftrightarrow> A<=C \<and> B<=D" |
13285 | 272 |
by blast |
273 |
||
76214 | 274 |
lemma qsum_equal_iff: "A <+> B = C <+> D \<longleftrightarrow> A=C \<and> B=D" |
13285 | 275 |
apply (simp (no_asm) add: extension qsum_subset_iff) |
276 |
apply blast |
|
277 |
done |
|
278 |
||
60770 | 279 |
subsubsection\<open>Eliminator -- qcase\<close> |
13285 | 280 |
|
281 |
lemma qcase_QInl [simp]: "qcase(c, d, QInl(a)) = c(a)" |
|
282 |
by (simp add: qsum_defs ) |
|
283 |
||
284 |
||
285 |
lemma qcase_QInr [simp]: "qcase(c, d, QInr(b)) = d(b)" |
|
286 |
by (simp add: qsum_defs ) |
|
287 |
||
288 |
lemma qcase_type: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
289 |
"\<lbrakk>u \<in> A <+> B; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
290 |
\<And>x. x \<in> A \<Longrightarrow> c(x): C(QInl(x)); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
291 |
\<And>y. y \<in> B \<Longrightarrow> d(y): C(QInr(y)) |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
292 |
\<rbrakk> \<Longrightarrow> qcase(c,d,u) \<in> C(u)" |
46953 | 293 |
by (simp add: qsum_defs, auto) |
13285 | 294 |
|
295 |
(** Rules for the Part primitive **) |
|
296 |
||
46953 | 297 |
lemma Part_QInl: "Part(A <+> B,QInl) = {QInl(x). x \<in> A}" |
13285 | 298 |
by blast |
299 |
||
46953 | 300 |
lemma Part_QInr: "Part(A <+> B,QInr) = {QInr(y). y \<in> B}" |
13285 | 301 |
by blast |
302 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
303 |
lemma Part_QInr2: "Part(A <+> B, \<lambda>x. QInr(h(x))) = {QInr(y). y \<in> Part(B,h)}" |
13285 | 304 |
by blast |
0 | 305 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
306 |
lemma Part_qsum_equality: "C \<subseteq> A <+> B \<Longrightarrow> Part(C,QInl) \<union> Part(C,QInr) = C" |
13285 | 307 |
by blast |
308 |
||
309 |
||
60770 | 310 |
subsubsection\<open>Monotonicity\<close> |
13285 | 311 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
312 |
lemma QPair_mono: "\<lbrakk>a<=c; b<=d\<rbrakk> \<Longrightarrow> <a;b> \<subseteq> <c;d>" |
13285 | 313 |
by (simp add: QPair_def sum_mono) |
314 |
||
315 |
lemma QSigma_mono [rule_format]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
316 |
"\<lbrakk>A<=C; \<forall>x\<in>A. B(x) \<subseteq> D(x)\<rbrakk> \<Longrightarrow> QSigma(A,B) \<subseteq> QSigma(C,D)" |
13285 | 317 |
by blast |
318 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
319 |
lemma QInl_mono: "a<=b \<Longrightarrow> QInl(a) \<subseteq> QInl(b)" |
13285 | 320 |
by (simp add: QInl_def subset_refl [THEN QPair_mono]) |
321 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
322 |
lemma QInr_mono: "a<=b \<Longrightarrow> QInr(a) \<subseteq> QInr(b)" |
13285 | 323 |
by (simp add: QInr_def subset_refl [THEN QPair_mono]) |
324 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
325 |
lemma qsum_mono: "\<lbrakk>A<=C; B<=D\<rbrakk> \<Longrightarrow> A <+> B \<subseteq> C <+> D" |
13285 | 326 |
by blast |
327 |
||
0 | 328 |
end |