author | haftmann |
Sun, 08 Oct 2017 22:28:22 +0200 | |
changeset 66814 | a24cde9588bb |
parent 62335 | e85c42f4f30a |
child 67091 | 1393c2340eec |
permissions | -rw-r--r-- |
55075 | 1 |
(* Title: HOL/Basic_BNFs.thy |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
3 |
Author: Andrei Popescu, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
Copyright 2012 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
|
49309
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
49236
diff
changeset
|
7 |
Registration of basic types as bounded natural functors. |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
|
60758 | 10 |
section \<open>Registration of Basic Types as Bounded Natural Functors\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
12 |
theory Basic_BNFs |
49310 | 13 |
imports BNF_Def |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
begin |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
15 |
|
58916 | 16 |
inductive_set setl :: "'a + 'b \<Rightarrow> 'a set" for s :: "'a + 'b" where |
17 |
"s = Inl x \<Longrightarrow> x \<in> setl s" |
|
18 |
inductive_set setr :: "'a + 'b \<Rightarrow> 'b set" for s :: "'a + 'b" where |
|
19 |
"s = Inr x \<Longrightarrow> x \<in> setr s" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
20 |
|
58916 | 21 |
lemma sum_set_defs[code]: |
22 |
"setl = (\<lambda>x. case x of Inl z => {z} | _ => {})" |
|
23 |
"setr = (\<lambda>x. case x of Inr z => {z} | _ => {})" |
|
24 |
by (auto simp: fun_eq_iff intro: setl.intros setr.intros elim: setl.cases setr.cases split: sum.splits) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
25 |
|
58916 | 26 |
lemma rel_sum_simps[code, simp]: |
55943 | 27 |
"rel_sum R1 R2 (Inl a1) (Inl b1) = R1 a1 b1" |
28 |
"rel_sum R1 R2 (Inl a1) (Inr b2) = False" |
|
29 |
"rel_sum R1 R2 (Inr a2) (Inl b1) = False" |
|
30 |
"rel_sum R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" |
|
58916 | 31 |
by (auto intro: rel_sum.intros elim: rel_sum.cases) |
55083 | 32 |
|
62324 | 33 |
inductive |
34 |
pred_sum :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> bool" for P1 P2 |
|
35 |
where |
|
36 |
"P1 a \<Longrightarrow> pred_sum P1 P2 (Inl a)" |
|
37 |
| "P2 b \<Longrightarrow> pred_sum P1 P2 (Inr b)" |
|
38 |
||
62335 | 39 |
lemma pred_sum_inject[code, simp]: |
40 |
"pred_sum P1 P2 (Inl a) \<longleftrightarrow> P1 a" |
|
41 |
"pred_sum P1 P2 (Inr b) \<longleftrightarrow> P2 b" |
|
42 |
by (simp add: pred_sum.simps)+ |
|
43 |
||
54421 | 44 |
bnf "'a + 'b" |
55931 | 45 |
map: map_sum |
54421 | 46 |
sets: setl setr |
47 |
bd: natLeq |
|
48 |
wits: Inl Inr |
|
55943 | 49 |
rel: rel_sum |
62324 | 50 |
pred: pred_sum |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
51 |
proof - |
55931 | 52 |
show "map_sum id id = id" by (rule map_sum.id) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
53 |
next |
54486 | 54 |
fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r" |
55931 | 55 |
show "map_sum (g1 o f1) (g2 o f2) = map_sum g1 g2 o map_sum f1 f2" |
56 |
by (rule map_sum.comp[symmetric]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
57 |
next |
54486 | 58 |
fix x and f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" and g1 g2 |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
59 |
assume a1: "\<And>z. z \<in> setl x \<Longrightarrow> f1 z = g1 z" and |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
60 |
a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z" |
55931 | 61 |
thus "map_sum f1 f2 x = map_sum g1 g2 x" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
62 |
proof (cases x) |
58916 | 63 |
case Inl thus ?thesis using a1 by (clarsimp simp: sum_set_defs(1)) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
64 |
next |
58916 | 65 |
case Inr thus ?thesis using a2 by (clarsimp simp: sum_set_defs(2)) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
66 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
67 |
next |
54486 | 68 |
fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 69 |
show "setl o map_sum f1 f2 = image f1 o setl" |
58916 | 70 |
by (rule ext, unfold o_apply) (simp add: sum_set_defs(1) split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
71 |
next |
54486 | 72 |
fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 73 |
show "setr o map_sum f1 f2 = image f2 o setr" |
58916 | 74 |
by (rule ext, unfold o_apply) (simp add: sum_set_defs(2) split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
75 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
76 |
show "card_order natLeq" by (rule natLeq_card_order) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
78 |
show "cinfinite natLeq" by (rule natLeq_cinfinite) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
79 |
next |
54486 | 80 |
fix x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
81 |
show "|setl x| \<le>o natLeq" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
82 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
83 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
58916 | 84 |
by (simp add: sum_set_defs(1) split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
85 |
next |
54486 | 86 |
fix x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
87 |
show "|setr x| \<le>o natLeq" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
88 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
89 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
58916 | 90 |
by (simp add: sum_set_defs(2) split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
91 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
92 |
fix R1 R2 S1 S2 |
55943 | 93 |
show "rel_sum R1 R2 OO rel_sum S1 S2 \<le> rel_sum (R1 OO S1) (R2 OO S2)" |
58916 | 94 |
by (force elim: rel_sum.cases) |
49453 | 95 |
next |
96 |
fix R S |
|
62324 | 97 |
show "rel_sum R S = (\<lambda>x y. |
98 |
\<exists>z. (setl z \<subseteq> {(x, y). R x y} \<and> setr z \<subseteq> {(x, y). S x y}) \<and> |
|
99 |
map_sum fst fst z = x \<and> map_sum snd snd z = y)" |
|
100 |
unfolding sum_set_defs relcompp.simps conversep.simps fun_eq_iff |
|
58916 | 101 |
by (fastforce elim: rel_sum.cases split: sum.splits) |
62324 | 102 |
qed (auto simp: sum_set_defs fun_eq_iff pred_sum.simps split: sum.splits) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
103 |
|
58916 | 104 |
inductive_set fsts :: "'a \<times> 'b \<Rightarrow> 'a set" for p :: "'a \<times> 'b" where |
105 |
"fst p \<in> fsts p" |
|
106 |
inductive_set snds :: "'a \<times> 'b \<Rightarrow> 'b set" for p :: "'a \<times> 'b" where |
|
107 |
"snd p \<in> snds p" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
108 |
|
58916 | 109 |
lemma prod_set_defs[code]: "fsts = (\<lambda>p. {fst p})" "snds = (\<lambda>p. {snd p})" |
110 |
by (auto intro: fsts.intros snds.intros elim: fsts.cases snds.cases) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
111 |
|
58916 | 112 |
inductive |
113 |
rel_prod :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b \<times> 'd \<Rightarrow> bool" for R1 R2 |
|
55083 | 114 |
where |
58916 | 115 |
"\<lbrakk>R1 a b; R2 c d\<rbrakk> \<Longrightarrow> rel_prod R1 R2 (a, c) (b, d)" |
116 |
||
62324 | 117 |
inductive |
118 |
pred_prod :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" for P1 P2 |
|
119 |
where |
|
120 |
"\<lbrakk>P1 a; P2 b\<rbrakk> \<Longrightarrow> pred_prod P1 P2 (a, b)" |
|
121 |
||
62335 | 122 |
lemma rel_prod_inject [code, simp]: |
58916 | 123 |
"rel_prod R1 R2 (a, b) (c, d) \<longleftrightarrow> R1 a c \<and> R2 b d" |
124 |
by (auto intro: rel_prod.intros elim: rel_prod.cases) |
|
125 |
||
62335 | 126 |
lemma pred_prod_inject [code, simp]: |
62324 | 127 |
"pred_prod P1 P2 (a, b) \<longleftrightarrow> P1 a \<and> P2 b" |
128 |
by (auto intro: pred_prod.intros elim: pred_prod.cases) |
|
129 |
||
58916 | 130 |
lemma rel_prod_conv: |
55944 | 131 |
"rel_prod R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)" |
58916 | 132 |
by (rule ext, rule ext) auto |
55083 | 133 |
|
62324 | 134 |
definition |
135 |
pred_fun :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" |
|
136 |
where |
|
137 |
"pred_fun A B = (\<lambda>f. \<forall>x. A x \<longrightarrow> B (f x))" |
|
138 |
||
139 |
lemma pred_funI: "(\<And>x. A x \<Longrightarrow> B (f x)) \<Longrightarrow> pred_fun A B f" |
|
140 |
unfolding pred_fun_def by simp |
|
141 |
||
54421 | 142 |
bnf "'a \<times> 'b" |
55932 | 143 |
map: map_prod |
54421 | 144 |
sets: fsts snds |
145 |
bd: natLeq |
|
55944 | 146 |
rel: rel_prod |
62324 | 147 |
pred: pred_prod |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
148 |
proof (unfold prod_set_defs) |
55932 | 149 |
show "map_prod id id = id" by (rule map_prod.id) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
150 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
151 |
fix f1 f2 g1 g2 |
55932 | 152 |
show "map_prod (g1 o f1) (g2 o f2) = map_prod g1 g2 o map_prod f1 f2" |
153 |
by (rule map_prod.comp[symmetric]) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
154 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
155 |
fix x f1 f2 g1 g2 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
156 |
assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z" |
55932 | 157 |
thus "map_prod f1 f2 x = map_prod g1 g2 x" by (cases x) simp |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
158 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
159 |
fix f1 f2 |
55932 | 160 |
show "(\<lambda>x. {fst x}) o map_prod f1 f2 = image f1 o (\<lambda>x. {fst x})" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
161 |
by (rule ext, unfold o_apply) simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
162 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
163 |
fix f1 f2 |
55932 | 164 |
show "(\<lambda>x. {snd x}) o map_prod f1 f2 = image f2 o (\<lambda>x. {snd x})" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
165 |
by (rule ext, unfold o_apply) simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
next |
52635
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents:
52545
diff
changeset
|
167 |
show "card_order natLeq" by (rule natLeq_card_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
168 |
next |
52635
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents:
52545
diff
changeset
|
169 |
show "cinfinite natLeq" by (rule natLeq_cinfinite) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
170 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
171 |
fix x |
52635
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents:
52545
diff
changeset
|
172 |
show "|{fst x}| \<le>o natLeq" |
55811 | 173 |
by (rule ordLess_imp_ordLeq) (simp add: finite_iff_ordLess_natLeq[symmetric]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
174 |
next |
52635
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents:
52545
diff
changeset
|
175 |
fix x |
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents:
52545
diff
changeset
|
176 |
show "|{snd x}| \<le>o natLeq" |
55811 | 177 |
by (rule ordLess_imp_ordLeq) (simp add: finite_iff_ordLess_natLeq[symmetric]) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
179 |
fix R1 R2 S1 S2 |
55944 | 180 |
show "rel_prod R1 R2 OO rel_prod S1 S2 \<le> rel_prod (R1 OO S1) (R2 OO S2)" by auto |
49453 | 181 |
next |
182 |
fix R S |
|
62324 | 183 |
show "rel_prod R S = (\<lambda>x y. |
184 |
\<exists>z. ({fst z} \<subseteq> {(x, y). R x y} \<and> {snd z} \<subseteq> {(x, y). S x y}) \<and> |
|
185 |
map_prod fst fst z = x \<and> map_prod snd snd z = y)" |
|
62335 | 186 |
unfolding prod_set_defs rel_prod_inject relcompp.simps conversep.simps fun_eq_iff |
49453 | 187 |
by auto |
62324 | 188 |
qed auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
189 |
|
54421 | 190 |
bnf "'a \<Rightarrow> 'b" |
191 |
map: "op \<circ>" |
|
192 |
sets: range |
|
193 |
bd: "natLeq +c |UNIV :: 'a set|" |
|
55945 | 194 |
rel: "rel_fun op =" |
62324 | 195 |
pred: "pred_fun (\<lambda>_. True)" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
196 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
197 |
fix f show "id \<circ> f = id f" by simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
198 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
199 |
fix f g show "op \<circ> (g \<circ> f) = op \<circ> g \<circ> op \<circ> f" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
200 |
unfolding comp_def[abs_def] .. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
201 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
202 |
fix x f g |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
203 |
assume "\<And>z. z \<in> range x \<Longrightarrow> f z = g z" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
204 |
thus "f \<circ> x = g \<circ> x" by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
205 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
206 |
fix f show "range \<circ> op \<circ> f = op ` f \<circ> range" |
56077 | 207 |
by (auto simp add: fun_eq_iff) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
208 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
209 |
show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)") |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
210 |
apply (rule card_order_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
211 |
apply (rule natLeq_card_order) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
212 |
by (rule card_of_card_order_on) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
213 |
(* *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
214 |
show "cinfinite (natLeq +c ?U)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
215 |
apply (rule cinfinite_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
216 |
apply (rule disjI1) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
217 |
by (rule natLeq_cinfinite) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
218 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
219 |
fix f :: "'d => 'a" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
220 |
have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image) |
54486 | 221 |
also have "?U \<le>o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
222 |
finally show "|range f| \<le>o natLeq +c ?U" . |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
223 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
224 |
fix R S |
55945 | 225 |
show "rel_fun op = R OO rel_fun op = S \<le> rel_fun op = (R OO S)" by (auto simp: rel_fun_def) |
49453 | 226 |
next |
49463 | 227 |
fix R |
62324 | 228 |
show "rel_fun op = R = (\<lambda>x y. |
229 |
\<exists>z. range z \<subseteq> {(x, y). R x y} \<and> fst \<circ> z = x \<and> snd \<circ> z = y)" |
|
230 |
unfolding rel_fun_def subset_iff by (force simp: fun_eq_iff[symmetric]) |
|
231 |
qed (auto simp: pred_fun_def) |
|
54191 | 232 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
233 |
end |