author | wenzelm |
Thu, 30 Jan 2025 22:29:45 +0100 | |
changeset 82023 | 9601f5582f33 |
parent 78099 | 4d9349989d94 |
permissions | -rw-r--r-- |
41777 | 1 |
(* Title: ZF/pair.thy |
13240 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1992 University of Cambridge |
|
4 |
*) |
|
5 |
||
60770 | 6 |
section\<open>Ordered Pairs\<close> |
13357 | 7 |
|
16417 | 8 |
theory pair imports upair |
42455 | 9 |
begin |
10 |
||
69605 | 11 |
ML_file \<open>simpdata.ML\<close> |
48891 | 12 |
|
60770 | 13 |
setup \<open> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48891
diff
changeset
|
14 |
map_theory_simpset |
60822 | 15 |
(Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt) |
45625
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
45620
diff
changeset
|
16 |
#> Simplifier.add_cong @{thm if_weak_cong}) |
60770 | 17 |
\<close> |
42794 | 18 |
|
69593 | 19 |
ML \<open>val ZF_ss = simpset_of \<^context>\<close> |
42794 | 20 |
|
76214 | 21 |
simproc_setup defined_Bex ("\<exists>x\<in>A. P(x) \<and> Q(x)") = \<open> |
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
76216
diff
changeset
|
22 |
K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def})) |
60770 | 23 |
\<close> |
42455 | 24 |
|
60770 | 25 |
simproc_setup defined_Ball ("\<forall>x\<in>A. P(x) \<longrightarrow> Q(x)") = \<open> |
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
76216
diff
changeset
|
26 |
K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def})) |
60770 | 27 |
\<close> |
42455 | 28 |
|
13240 | 29 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
30 |
(** Lemmas for showing that \<langle>a,b\<rangle> uniquely determines a and b **) |
13240 | 31 |
|
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
32 |
lemma singleton_eq_iff [iff]: "{a} = {b} \<longleftrightarrow> a=b" |
13240 | 33 |
by (rule extension [THEN iff_trans], blast) |
34 |
||
76214 | 35 |
lemma doubleton_eq_iff: "{a,b} = {c,d} \<longleftrightarrow> (a=c \<and> b=d) | (a=d \<and> b=c)" |
13240 | 36 |
by (rule extension [THEN iff_trans], blast) |
37 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
38 |
lemma Pair_iff [simp]: "\<langle>a,b\<rangle> = \<langle>c,d\<rangle> \<longleftrightarrow> a=c \<and> b=d" |
13240 | 39 |
by (simp add: Pair_def doubleton_eq_iff, blast) |
40 |
||
45602 | 41 |
lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!] |
13240 | 42 |
|
45602 | 43 |
lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1] |
44 |
lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2] |
|
13240 | 45 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
46 |
lemma Pair_not_0: "\<langle>a,b\<rangle> \<noteq> 0" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
47 |
unfolding Pair_def |
13240 | 48 |
apply (blast elim: equalityE) |
49 |
done |
|
50 |
||
45602 | 51 |
lemmas Pair_neq_0 = Pair_not_0 [THEN notE, elim!] |
13240 | 52 |
|
53 |
declare sym [THEN Pair_neq_0, elim!] |
|
54 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
55 |
lemma Pair_neq_fst: "\<langle>a,b\<rangle>=a \<Longrightarrow> P" |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
56 |
proof (unfold Pair_def) |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
57 |
assume eq: "{{a, a}, {a, b}} = a" |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
58 |
have "{a, a} \<in> {{a, a}, {a, b}}" by (rule consI1) |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
59 |
hence "{a, a} \<in> a" by (simp add: eq) |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
60 |
moreover have "a \<in> {a, a}" by (rule consI1) |
46953 | 61 |
ultimately show "P" by (rule mem_asym) |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
62 |
qed |
13240 | 63 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
64 |
lemma Pair_neq_snd: "\<langle>a,b\<rangle>=b \<Longrightarrow> P" |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
65 |
proof (unfold Pair_def) |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
66 |
assume eq: "{{a, a}, {a, b}} = b" |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
67 |
have "{a, b} \<in> {{a, a}, {a, b}}" by blast |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
68 |
hence "{a, b} \<in> b" by (simp add: eq) |
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
69 |
moreover have "b \<in> {a, b}" by blast |
46953 | 70 |
ultimately show "P" by (rule mem_asym) |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
71 |
qed |
13240 | 72 |
|
73 |
||
60770 | 74 |
subsection\<open>Sigma: Disjoint Union of a Family of Sets\<close> |
13357 | 75 |
|
60770 | 76 |
text\<open>Generalizes Cartesian product\<close> |
13240 | 77 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
78 |
lemma Sigma_iff [simp]: "\<langle>a,b\<rangle>: Sigma(A,B) \<longleftrightarrow> a \<in> A \<and> b \<in> B(a)" |
13240 | 79 |
by (simp add: Sigma_def) |
80 |
||
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
81 |
lemma SigmaI [TC,intro!]: "\<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Sigma(A,B)" |
13240 | 82 |
by simp |
83 |
||
45602 | 84 |
lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1] |
85 |
lemmas SigmaD2 = Sigma_iff [THEN iffD1, THEN conjunct2] |
|
13240 | 86 |
|
87 |
(*The general elimination rule*) |
|
88 |
lemma SigmaE [elim!]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
89 |
"\<lbrakk>c \<in> Sigma(A,B); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
90 |
\<And>x y.\<lbrakk>x \<in> A; y \<in> B(x); c=\<langle>x,y\<rangle>\<rbrakk> \<Longrightarrow> P |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
91 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 92 |
by (unfold Sigma_def, blast) |
13240 | 93 |
|
94 |
lemma SigmaE2 [elim!]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
95 |
"\<lbrakk>\<langle>a,b\<rangle> \<in> Sigma(A,B); |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
96 |
\<lbrakk>a \<in> A; b \<in> B(a)\<rbrakk> \<Longrightarrow> P |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
97 |
\<rbrakk> \<Longrightarrow> P" |
46953 | 98 |
by (unfold Sigma_def, blast) |
13240 | 99 |
|
100 |
lemma Sigma_cong: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
101 |
"\<lbrakk>A=A'; \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow> |
13240 | 102 |
Sigma(A,B) = Sigma(A',B')" |
103 |
by (simp add: Sigma_def) |
|
104 |
||
105 |
(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause |
|
106 |
flex-flex pairs and the "Check your prover" error. Most |
|
107 |
Sigmas and Pis are abbreviated as * or -> *) |
|
108 |
||
109 |
lemma Sigma_empty1 [simp]: "Sigma(0,B) = 0" |
|
110 |
by blast |
|
111 |
||
112 |
lemma Sigma_empty2 [simp]: "A*0 = 0" |
|
113 |
by blast |
|
114 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
115 |
lemma Sigma_empty_iff: "A*B=0 \<longleftrightarrow> A=0 | B=0" |
13240 | 116 |
by blast |
117 |
||
118 |
||
69593 | 119 |
subsection\<open>Projections \<^term>\<open>fst\<close> and \<^term>\<open>snd\<close>\<close> |
13240 | 120 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
121 |
lemma fst_conv [simp]: "fst(\<langle>a,b\<rangle>) = a" |
13544 | 122 |
by (simp add: fst_def) |
13240 | 123 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
124 |
lemma snd_conv [simp]: "snd(\<langle>a,b\<rangle>) = b" |
13544 | 125 |
by (simp add: snd_def) |
13240 | 126 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
127 |
lemma fst_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> fst(p) \<in> A" |
13240 | 128 |
by auto |
129 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
130 |
lemma snd_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> snd(p) \<in> B(fst(p))" |
13240 | 131 |
by auto |
132 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
133 |
lemma Pair_fst_snd_eq: "a \<in> Sigma(A,B) \<Longrightarrow> <fst(a),snd(a)> = a" |
13240 | 134 |
by auto |
135 |
||
136 |
||
69593 | 137 |
subsection\<open>The Eliminator, \<^term>\<open>split\<close>\<close> |
13240 | 138 |
|
139 |
(*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
|
140 |
lemma split [simp]: "split(\<lambda>x y. c(x,y), \<langle>a,b\<rangle>) \<equiv> c(a,b)" |
13240 | 141 |
by (simp add: split_def) |
142 |
||
143 |
lemma split_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
144 |
"\<lbrakk>p \<in> Sigma(A,B); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
145 |
\<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(\<langle>x,y\<rangle>) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
146 |
\<rbrakk> \<Longrightarrow> split(\<lambda>x y. c(x,y), p) \<in> C(p)" |
46953 | 147 |
by (erule SigmaE, auto) |
13240 | 148 |
|
46953 | 149 |
lemma expand_split: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
150 |
"u \<in> A*B \<Longrightarrow> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
151 |
R(split(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = \<langle>x,y\<rangle> \<longrightarrow> R(c(x,y)))" |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
152 |
by (auto simp add: split_def) |
13240 | 153 |
|
154 |
||
69593 | 155 |
subsection\<open>A version of \<^term>\<open>split\<close> for Formulae: Result Type \<^typ>\<open>o\<close>\<close> |
13240 | 156 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
157 |
lemma splitI: "R(a,b) \<Longrightarrow> split(R, \<langle>a,b\<rangle>)" |
13240 | 158 |
by (simp add: split_def) |
159 |
||
160 |
lemma splitE: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
161 |
"\<lbrakk>split(R,z); z \<in> Sigma(A,B); |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
162 |
\<And>x y. \<lbrakk>z = \<langle>x,y\<rangle>; R(x,y)\<rbrakk> \<Longrightarrow> P |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71886
diff
changeset
|
163 |
\<rbrakk> \<Longrightarrow> P" |
46841
49b91b716cbe
Structured and calculation-based proofs (with new trans rules!)
paulson
parents:
46821
diff
changeset
|
164 |
by (auto simp add: split_def) |
13240 | 165 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
166 |
lemma splitD: "split(R,\<langle>a,b\<rangle>) \<Longrightarrow> R(a,b)" |
13240 | 167 |
by (simp add: split_def) |
168 |
||
60770 | 169 |
text \<open> |
14864 | 170 |
\bigskip Complex rules for Sigma. |
60770 | 171 |
\<close> |
14864 | 172 |
|
173 |
lemma split_paired_Bex_Sigma [simp]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
174 |
"(\<exists>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<exists>x \<in> A. \<exists>y \<in> B(x). P(\<langle>x,y\<rangle>))" |
14864 | 175 |
by blast |
176 |
||
177 |
lemma split_paired_Ball_Sigma [simp]: |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
178 |
"(\<forall>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<forall>x \<in> A. \<forall>y \<in> B(x). P(\<langle>x,y\<rangle>))" |
14864 | 179 |
by blast |
180 |
||
9570
e16e168984e1
installation of cancellation simprocs for the integers
paulson
parents:
2469
diff
changeset
|
181 |
end |
124 | 182 |
|
2469 | 183 |