author | haftmann |
Mon, 20 Oct 2014 07:45:58 +0200 | |
changeset 58710 | 7216a10d69ba |
parent 58446 | e89f57d1e46c |
child 58889 | 5b7a9633cfa8 |
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 |
|
49309
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
49236
diff
changeset
|
10 |
header {* 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
|
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 |
|
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
16 |
definition setl :: "'a + 'b \<Rightarrow> 'a set" where |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
17 |
"setl x = (case x of Inl z => {z} | _ => {})" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
18 |
|
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
19 |
definition setr :: "'a + 'b \<Rightarrow> 'b set" where |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
20 |
"setr x = (case x of Inr z => {z} | _ => {})" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
21 |
|
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
22 |
lemmas sum_set_defs = setl_def[abs_def] setr_def[abs_def] |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
23 |
|
55943 | 24 |
lemma rel_sum_simps[simp]: |
25 |
"rel_sum R1 R2 (Inl a1) (Inl b1) = R1 a1 b1" |
|
26 |
"rel_sum R1 R2 (Inl a1) (Inr b2) = False" |
|
27 |
"rel_sum R1 R2 (Inr a2) (Inl b1) = False" |
|
28 |
"rel_sum R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" |
|
29 |
unfolding rel_sum_def by simp_all |
|
55083 | 30 |
|
54421 | 31 |
bnf "'a + 'b" |
55931 | 32 |
map: map_sum |
54421 | 33 |
sets: setl setr |
34 |
bd: natLeq |
|
35 |
wits: Inl Inr |
|
55943 | 36 |
rel: rel_sum |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
37 |
proof - |
55931 | 38 |
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
|
39 |
next |
54486 | 40 |
fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r" |
55931 | 41 |
show "map_sum (g1 o f1) (g2 o f2) = map_sum g1 g2 o map_sum f1 f2" |
42 |
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
|
43 |
next |
54486 | 44 |
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
|
45 |
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
|
46 |
a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z" |
55931 | 47 |
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
|
48 |
proof (cases x) |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
49 |
case Inl thus ?thesis using a1 by (clarsimp simp: setl_def) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
50 |
next |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
51 |
case Inr thus ?thesis using a2 by (clarsimp simp: setr_def) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
52 |
qed |
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> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 55 |
show "setl o map_sum f1 f2 = image f1 o setl" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
56 |
by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split) |
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 f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 59 |
show "setr o map_sum f1 f2 = image f2 o setr" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
60 |
by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
61 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
62 |
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
|
63 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
64 |
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
|
65 |
next |
54486 | 66 |
fix x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
67 |
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
|
68 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
69 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
70 |
by (simp add: setl_def 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 x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
73 |
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
|
74 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
75 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
76 |
by (simp add: setr_def split: sum.split) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
78 |
fix R1 R2 S1 S2 |
55943 | 79 |
show "rel_sum R1 R2 OO rel_sum S1 S2 \<le> rel_sum (R1 OO S1) (R2 OO S2)" |
80 |
by (auto simp: rel_sum_def split: sum.splits) |
|
49453 | 81 |
next |
82 |
fix R S |
|
55943 | 83 |
show "rel_sum R S = |
55931 | 84 |
(Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (map_sum fst fst))\<inverse>\<inverse> OO |
85 |
Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (map_sum snd snd)" |
|
55943 | 86 |
unfolding setl_def setr_def rel_sum_def Grp_def relcompp.simps conversep.simps fun_eq_iff |
49453 | 87 |
by (fastforce split: sum.splits) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
88 |
qed (auto simp: sum_set_defs) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
89 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
90 |
definition fsts :: "'a \<times> 'b \<Rightarrow> 'a set" where |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
91 |
"fsts x = {fst x}" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
92 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
93 |
definition snds :: "'a \<times> 'b \<Rightarrow> 'b set" where |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
94 |
"snds x = {snd x}" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
95 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
96 |
lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def] |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
97 |
|
55083 | 98 |
definition |
55944 | 99 |
rel_prod :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b \<times> 'd \<Rightarrow> bool" |
55083 | 100 |
where |
55944 | 101 |
"rel_prod R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)" |
55083 | 102 |
|
55944 | 103 |
lemma rel_prod_apply [simp]: |
104 |
"rel_prod R1 R2 (a, b) (c, d) \<longleftrightarrow> R1 a c \<and> R2 b d" |
|
105 |
by (simp add: rel_prod_def) |
|
55083 | 106 |
|
54421 | 107 |
bnf "'a \<times> 'b" |
55932 | 108 |
map: map_prod |
54421 | 109 |
sets: fsts snds |
110 |
bd: natLeq |
|
55944 | 111 |
rel: rel_prod |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
112 |
proof (unfold prod_set_defs) |
55932 | 113 |
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
|
114 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
115 |
fix f1 f2 g1 g2 |
55932 | 116 |
show "map_prod (g1 o f1) (g2 o f2) = map_prod g1 g2 o map_prod f1 f2" |
117 |
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
|
118 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
119 |
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
|
120 |
assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z" |
55932 | 121 |
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
|
122 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
123 |
fix f1 f2 |
55932 | 124 |
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
|
125 |
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
|
126 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
127 |
fix f1 f2 |
55932 | 128 |
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
|
129 |
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
|
130 |
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
|
131 |
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
|
132 |
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
|
133 |
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
|
134 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
135 |
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
|
136 |
show "|{fst x}| \<le>o natLeq" |
55811 | 137 |
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
|
138 |
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
|
139 |
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
|
140 |
show "|{snd x}| \<le>o natLeq" |
55811 | 141 |
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
|
142 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
143 |
fix R1 R2 S1 S2 |
55944 | 144 |
show "rel_prod R1 R2 OO rel_prod S1 S2 \<le> rel_prod (R1 OO S1) (R2 OO S2)" by auto |
49453 | 145 |
next |
146 |
fix R S |
|
55944 | 147 |
show "rel_prod R S = |
55932 | 148 |
(Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_prod fst fst))\<inverse>\<inverse> OO |
149 |
Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_prod snd snd)" |
|
55944 | 150 |
unfolding prod_set_defs rel_prod_def Grp_def relcompp.simps conversep.simps fun_eq_iff |
49453 | 151 |
by auto |
54189
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
traytel
parents:
53026
diff
changeset
|
152 |
qed |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
153 |
|
54421 | 154 |
bnf "'a \<Rightarrow> 'b" |
155 |
map: "op \<circ>" |
|
156 |
sets: range |
|
157 |
bd: "natLeq +c |UNIV :: 'a set|" |
|
55945 | 158 |
rel: "rel_fun op =" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
159 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
160 |
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
|
161 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
162 |
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
|
163 |
unfolding comp_def[abs_def] .. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
164 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
165 |
fix x f g |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
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
|
167 |
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
|
168 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
169 |
fix f show "range \<circ> op \<circ> f = op ` f \<circ> range" |
56077 | 170 |
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
|
171 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
172 |
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
|
173 |
apply (rule card_order_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
174 |
apply (rule natLeq_card_order) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
175 |
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
|
176 |
(* *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
177 |
show "cinfinite (natLeq +c ?U)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
apply (rule cinfinite_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
179 |
apply (rule disjI1) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
180 |
by (rule natLeq_cinfinite) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
181 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
182 |
fix f :: "'d => 'a" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
183 |
have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image) |
54486 | 184 |
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
|
185 |
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
|
186 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
187 |
fix R S |
55945 | 188 |
show "rel_fun op = R OO rel_fun op = S \<le> rel_fun op = (R OO S)" by (auto simp: rel_fun_def) |
49453 | 189 |
next |
49463 | 190 |
fix R |
55945 | 191 |
show "rel_fun op = R = |
51893
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51836
diff
changeset
|
192 |
(Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> fst))\<inverse>\<inverse> OO |
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents:
51836
diff
changeset
|
193 |
Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)" |
55945 | 194 |
unfolding rel_fun_def Grp_def fun_eq_iff relcompp.simps conversep.simps subset_iff image_iff |
55811 | 195 |
comp_apply mem_Collect_eq split_beta bex_UNIV |
196 |
proof (safe, unfold fun_eq_iff[symmetric]) |
|
197 |
fix x xa a b c xb y aa ba |
|
198 |
assume *: "x = a" "xa = c" "a = ba" "b = aa" "c = (\<lambda>x. snd (b x))" "ba = (\<lambda>x. fst (aa x))" and |
|
199 |
**: "\<forall>t. (\<exists>x. t = aa x) \<longrightarrow> R (fst t) (snd t)" |
|
200 |
show "R (x y) (xa y)" unfolding * by (rule mp[OF spec[OF **]]) blast |
|
201 |
qed force |
|
54189
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
traytel
parents:
53026
diff
changeset
|
202 |
qed |
54191 | 203 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
204 |
end |