author | wenzelm |
Fri, 24 Feb 2023 11:07:31 +0100 | |
changeset 77363 | cbd053fff24c |
parent 76217 | 8655344f1cf6 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/Coind/Map.thy |
2 |
Author: Jacob Frost, Cambridge University Computer Laboratory |
|
915 | 3 |
Copyright 1995 University of Cambridge |
12595 | 4 |
|
35762 | 5 |
Some sample proofs of inclusions for the final coalgebra "U" (by lcp). |
915 | 6 |
*) |
7 |
||
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
61980
diff
changeset
|
8 |
theory Map imports ZF begin |
915 | 9 |
|
24893 | 10 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
11 |
TMap :: "[i,i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
12 |
"TMap(A,B) \<equiv> {m \<in> Pow(A*\<Union>(B)).\<forall>a \<in> A. m``{a} \<in> B}" |
11318 | 13 |
|
24893 | 14 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
15 |
PMap :: "[i,i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
16 |
"PMap(A,B) \<equiv> TMap(A,cons(0,B))" |
915 | 17 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
18 |
(* Note: 0 \<in> B \<Longrightarrow> TMap(A,B) = PMap(A,B) *) |
915 | 19 |
|
24893 | 20 |
definition |
21 |
map_emp :: i where |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
22 |
"map_emp \<equiv> 0" |
11318 | 23 |
|
24893 | 24 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
25 |
map_owr :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
26 |
"map_owr(m,a,b) \<equiv> \<Sum>x \<in> {a} \<union> domain(m). if x=a then b else m``{x}" |
12595 | 27 |
|
24893 | 28 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
29 |
map_app :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
30 |
"map_app(m,a) \<equiv> m``{a}" |
915 | 31 |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
32 |
lemma "{0,1} \<subseteq> {1} \<union> TMap(I, {0,1})" |
12595 | 33 |
by (unfold TMap_def, blast) |
34 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
35 |
lemma "{0} \<union> TMap(I,1) \<subseteq> {1} \<union> TMap(I, {0} \<union> TMap(I,1))" |
12595 | 36 |
by (unfold TMap_def, blast) |
37 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
38 |
lemma "{0,1} \<union> TMap(I,2) \<subseteq> {1} \<union> TMap(I, {0,1} \<union> TMap(I,2))" |
12595 | 39 |
by (unfold TMap_def, blast) |
40 |
||
41 |
(*A bit too slow. |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
42 |
lemma "{0,1} \<union> TMap(I,TMap(I,2)) \<union> TMap(I,2) \<subseteq> |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35762
diff
changeset
|
43 |
{1} \<union> TMap(I, {0,1} \<union> TMap(I,TMap(I,2)) \<union> TMap(I,2))" |
12595 | 44 |
by (unfold TMap_def, blast) |
45 |
*) |
|
46 |
||
47 |
(* ############################################################ *) |
|
48 |
(* Lemmas *) |
|
49 |
(* ############################################################ *) |
|
50 |
||
51 |
lemma qbeta_if: "Sigma(A,B)``{a} = (if a \<in> A then B(a) else 0)" |
|
52 |
by auto |
|
53 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
54 |
lemma qbeta: "a \<in> A \<Longrightarrow> Sigma(A,B)``{a} = B(a)" |
12595 | 55 |
by fast |
56 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
57 |
lemma qbeta_emp: "a\<notin>A \<Longrightarrow> Sigma(A,B)``{a} = 0" |
12595 | 58 |
by fast |
59 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
60 |
lemma image_Sigma1: "a \<notin> A \<Longrightarrow> Sigma(A,B)``{a}=0" |
12595 | 61 |
by fast |
62 |
||
63 |
||
64 |
(* ############################################################ *) |
|
65 |
(* Inclusion in Quine Universes *) |
|
66 |
(* ############################################################ *) |
|
67 |
||
68 |
(* Lemmas *) |
|
69 |
||
70 |
lemma MapQU_lemma: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
71 |
"A \<subseteq> univ(X) \<Longrightarrow> Pow(A * \<Union>(quniv(X))) \<subseteq> quniv(X)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
72 |
unfolding quniv_def |
12595 | 73 |
apply (rule Pow_mono) |
74 |
apply (rule subset_trans [OF Sigma_mono product_univ]) |
|
75 |
apply (erule subset_trans) |
|
76 |
apply (rule arg_subset_eclose [THEN univ_mono]) |
|
77 |
apply (simp add: Union_Pow_eq) |
|
78 |
done |
|
79 |
||
80 |
(* Theorems *) |
|
81 |
||
82 |
lemma mapQU: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
83 |
"\<lbrakk>m \<in> PMap(A,quniv(B)); \<And>x. x \<in> A \<Longrightarrow> x \<in> univ(B)\<rbrakk> \<Longrightarrow> m \<in> quniv(B)" |
76217 | 84 |
unfolding PMap_def TMap_def |
12595 | 85 |
apply (blast intro!: MapQU_lemma [THEN subsetD]) |
86 |
done |
|
87 |
||
88 |
(* ############################################################ *) |
|
89 |
(* Monotonicity *) |
|
90 |
(* ############################################################ *) |
|
91 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
92 |
lemma PMap_mono: "B \<subseteq> C \<Longrightarrow> PMap(A,B)<=PMap(A,C)" |
12595 | 93 |
by (unfold PMap_def TMap_def, blast) |
94 |
||
95 |
||
96 |
(* ############################################################ *) |
|
97 |
(* Introduction Rules *) |
|
98 |
(* ############################################################ *) |
|
99 |
||
100 |
(** map_emp **) |
|
101 |
||
102 |
lemma pmap_empI: "map_emp \<in> PMap(A,B)" |
|
103 |
by (unfold map_emp_def PMap_def TMap_def, auto) |
|
104 |
||
105 |
(** map_owr **) |
|
106 |
||
107 |
lemma pmap_owrI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
108 |
"\<lbrakk>m \<in> PMap(A,B); a \<in> A; b \<in> B\<rbrakk> \<Longrightarrow> map_owr(m,a,b):PMap(A,B)" |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12595
diff
changeset
|
109 |
apply (unfold map_owr_def PMap_def TMap_def, safe) |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12595
diff
changeset
|
110 |
apply (simp_all add: if_iff, auto) |
12595 | 111 |
(*Remaining subgoal*) |
112 |
apply (rule excluded_middle [THEN disjE]) |
|
113 |
apply (erule image_Sigma1) |
|
59788 | 114 |
apply (drule_tac psi = "uu \<notin> B" for uu in asm_rl) |
12595 | 115 |
apply (auto simp add: qbeta) |
116 |
done |
|
117 |
||
118 |
(** map_app **) |
|
119 |
||
120 |
lemma tmap_app_notempty: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
121 |
"\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a) \<noteq>0" |
12595 | 122 |
by (unfold TMap_def map_app_def, blast) |
123 |
||
124 |
lemma tmap_appI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
125 |
"\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a):B" |
12595 | 126 |
by (unfold TMap_def map_app_def domain_def, blast) |
127 |
||
128 |
lemma pmap_appI: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
129 |
"\<lbrakk>m \<in> PMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> map_app(m,a):B" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
130 |
unfolding PMap_def |
12595 | 131 |
apply (frule tmap_app_notempty, assumption) |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12595
diff
changeset
|
132 |
apply (drule tmap_appI, auto) |
12595 | 133 |
done |
134 |
||
135 |
(** domain **) |
|
136 |
||
137 |
lemma tmap_domainD: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
138 |
"\<lbrakk>m \<in> TMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> a \<in> A" |
12595 | 139 |
by (unfold TMap_def, blast) |
140 |
||
141 |
lemma pmap_domainD: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
142 |
"\<lbrakk>m \<in> PMap(A,B); a \<in> domain(m)\<rbrakk> \<Longrightarrow> a \<in> A" |
12595 | 143 |
by (unfold PMap_def TMap_def, blast) |
144 |
||
145 |
||
146 |
(* ############################################################ *) |
|
147 |
(* Equalities *) |
|
148 |
(* ############################################################ *) |
|
149 |
||
150 |
(** Domain **) |
|
151 |
||
152 |
(* Lemmas *) |
|
153 |
||
154 |
lemma domain_UN: "domain(\<Union>x \<in> A. B(x)) = (\<Union>x \<in> A. domain(B(x)))" |
|
155 |
by fast |
|
156 |
||
157 |
||
158 |
lemma domain_Sigma: "domain(Sigma(A,B)) = {x \<in> A. \<exists>y. y \<in> B(x)}" |
|
159 |
by blast |
|
160 |
||
161 |
(* Theorems *) |
|
162 |
||
163 |
lemma map_domain_emp: "domain(map_emp) = 0" |
|
164 |
by (unfold map_emp_def, blast) |
|
165 |
||
166 |
lemma map_domain_owr: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
167 |
"b \<noteq> 0 \<Longrightarrow> domain(map_owr(f,a,b)) = {a} \<union> domain(f)" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
168 |
unfolding map_owr_def |
12595 | 169 |
apply (auto simp add: domain_Sigma) |
170 |
done |
|
171 |
||
172 |
(** Application **) |
|
173 |
||
174 |
lemma map_app_owr: |
|
175 |
"map_app(map_owr(f,a,b),c) = (if c=a then b else map_app(f,c))" |
|
176 |
by (simp add: qbeta_if map_app_def map_owr_def, blast) |
|
177 |
||
178 |
lemma map_app_owr1: "map_app(map_owr(f,a,b),a) = b" |
|
179 |
by (simp add: map_app_owr) |
|
180 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
65449
diff
changeset
|
181 |
lemma map_app_owr2: "c \<noteq> a \<Longrightarrow> map_app(map_owr(f,a,b),c)= map_app(f,c)" |
12595 | 182 |
by (simp add: map_app_owr) |
183 |
||
915 | 184 |
end |