| author | desharna | 
| Mon, 01 Sep 2014 13:23:39 +0200 | |
| changeset 58104 | c5316f843f72 | 
| parent 46822 | 95f1e700b712 | 
| child 59788 | 6f7b6adac439 | 
| 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  | 
||
| 16417 | 8  | 
theory Map imports Main begin  | 
| 915 | 9  | 
|
| 24893 | 10  | 
definition  | 
11  | 
TMap :: "[i,i] => i" where  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
35762 
diff
changeset
 | 
12  | 
   "TMap(A,B) == {m \<in> Pow(A*\<Union>(B)).\<forall>a \<in> A. m``{a} \<in> B}"
 | 
| 11318 | 13  | 
|
| 24893 | 14  | 
definition  | 
15  | 
PMap :: "[i,i] => i" where  | 
|
| 11318 | 16  | 
"PMap(A,B) == TMap(A,cons(0,B))"  | 
| 915 | 17  | 
|
| 12595 | 18  | 
(* Note: 0 \<in> B ==> TMap(A,B) = PMap(A,B) *)  | 
| 915 | 19  | 
|
| 24893 | 20  | 
definition  | 
21  | 
map_emp :: i where  | 
|
| 11318 | 22  | 
"map_emp == 0"  | 
23  | 
||
| 24893 | 24  | 
definition  | 
25  | 
map_owr :: "[i,i,i]=>i" where  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
35762 
diff
changeset
 | 
26  | 
   "map_owr(m,a,b) == \<Sigma> x \<in> {a} \<union> domain(m). if x=a then b else m``{x}"
 | 
| 12595 | 27  | 
|
| 24893 | 28  | 
definition  | 
29  | 
map_app :: "[i,i]=>i" where  | 
|
| 11318 | 30  | 
   "map_app(m,a) == 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  | 
||
54  | 
lemma qbeta: "a \<in> A ==> Sigma(A,B)``{a} = B(a)"
 | 
|
55  | 
by fast  | 
|
56  | 
||
57  | 
lemma qbeta_emp: "a\<notin>A ==> Sigma(A,B)``{a} = 0"
 | 
|
58  | 
by fast  | 
|
59  | 
||
60  | 
lemma image_Sigma1: "a \<notin> A ==> Sigma(A,B)``{a}=0"
 | 
|
61  | 
by fast  | 
|
62  | 
||
63  | 
||
64  | 
(* ############################################################ *)  | 
|
65  | 
(* Inclusion in Quine Universes *)  | 
|
66  | 
(* ############################################################ *)  | 
|
67  | 
||
68  | 
(* Lemmas *)  | 
|
69  | 
||
70  | 
lemma MapQU_lemma:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
35762 
diff
changeset
 | 
71  | 
"A \<subseteq> univ(X) ==> Pow(A * \<Union>(quniv(X))) \<subseteq> quniv(X)"  | 
| 12595 | 72  | 
apply (unfold quniv_def)  | 
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:  | 
|
83  | 
"[| m \<in> PMap(A,quniv(B)); !!x. x \<in> A ==> x \<in> univ(B) |] ==> m \<in> quniv(B)"  | 
|
84  | 
apply (unfold PMap_def TMap_def)  | 
|
85  | 
apply (blast intro!: MapQU_lemma [THEN subsetD])  | 
|
86  | 
done  | 
|
87  | 
||
88  | 
(* ############################################################ *)  | 
|
89  | 
(* Monotonicity *)  | 
|
90  | 
(* ############################################################ *)  | 
|
91  | 
||
92  | 
lemma PMap_mono: "B \<subseteq> C ==> PMap(A,B)<=PMap(A,C)"  | 
|
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:  | 
|
108  | 
"[| m \<in> PMap(A,B); a \<in> A; b \<in> B |] ==> 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)  | 
|
114  | 
apply (drule_tac psi = "?uu \<notin> B" in asm_rl)  | 
|
115  | 
apply (auto simp add: qbeta)  | 
|
116  | 
done  | 
|
117  | 
||
118  | 
(** map_app **)  | 
|
119  | 
||
120  | 
lemma tmap_app_notempty:  | 
|
121  | 
"[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> map_app(m,a) \<noteq>0"  | 
|
122  | 
by (unfold TMap_def map_app_def, blast)  | 
|
123  | 
||
124  | 
lemma tmap_appI:  | 
|
125  | 
"[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> map_app(m,a):B"  | 
|
126  | 
by (unfold TMap_def map_app_def domain_def, blast)  | 
|
127  | 
||
128  | 
lemma pmap_appI:  | 
|
129  | 
"[| m \<in> PMap(A,B); a \<in> domain(m) |] ==> map_app(m,a):B"  | 
|
130  | 
apply (unfold PMap_def)  | 
|
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:  | 
|
138  | 
"[| m \<in> TMap(A,B); a \<in> domain(m) |] ==> a \<in> A"  | 
|
139  | 
by (unfold TMap_def, blast)  | 
|
140  | 
||
141  | 
lemma pmap_domainD:  | 
|
142  | 
"[| m \<in> PMap(A,B); a \<in> domain(m) |] ==> a \<in> A"  | 
|
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:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
35762 
diff
changeset
 | 
167  | 
  "b \<noteq> 0 ==> domain(map_owr(f,a,b)) = {a} \<union> domain(f)"
 | 
| 12595 | 168  | 
apply (unfold map_owr_def)  | 
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  | 
||
181  | 
lemma map_app_owr2: "c \<noteq> a ==> map_app(map_owr(f,a,b),c)= map_app(f,c)"  | 
|
182  | 
by (simp add: map_app_owr)  | 
|
183  | 
||
| 915 | 184  | 
end  |