author | wenzelm |
Mon, 28 Dec 2015 01:26:34 +0100 | |
changeset 61944 | 5d06ecfdb472 |
parent 61681 | ca53150406c9 |
child 62324 | ae44f16dcea5 |
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 |
|
54421 | 33 |
bnf "'a + 'b" |
55931 | 34 |
map: map_sum |
54421 | 35 |
sets: setl setr |
36 |
bd: natLeq |
|
37 |
wits: Inl Inr |
|
55943 | 38 |
rel: rel_sum |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
39 |
proof - |
55931 | 40 |
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
|
41 |
next |
54486 | 42 |
fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r" |
55931 | 43 |
show "map_sum (g1 o f1) (g2 o f2) = map_sum g1 g2 o map_sum f1 f2" |
44 |
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
|
45 |
next |
54486 | 46 |
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
|
47 |
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
|
48 |
a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z" |
55931 | 49 |
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
|
50 |
proof (cases x) |
58916 | 51 |
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
|
52 |
next |
58916 | 53 |
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
|
54 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
55 |
next |
54486 | 56 |
fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 57 |
show "setl o map_sum f1 f2 = image f1 o setl" |
58916 | 58 |
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
|
59 |
next |
54486 | 60 |
fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" |
55931 | 61 |
show "setr o map_sum f1 f2 = image f2 o setr" |
58916 | 62 |
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
|
63 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
64 |
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
|
65 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
66 |
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
|
67 |
next |
54486 | 68 |
fix x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
69 |
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
|
70 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
71 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
58916 | 72 |
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
|
73 |
next |
54486 | 74 |
fix x :: "'o + 'p" |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49450
diff
changeset
|
75 |
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
|
76 |
apply (rule ordLess_imp_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) |
58916 | 78 |
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
|
79 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
80 |
fix R1 R2 S1 S2 |
55943 | 81 |
show "rel_sum R1 R2 OO rel_sum S1 S2 \<le> rel_sum (R1 OO S1) (R2 OO S2)" |
58916 | 82 |
by (force elim: rel_sum.cases) |
49453 | 83 |
next |
84 |
fix R S |
|
55943 | 85 |
show "rel_sum R S = |
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
86 |
(Grp {x. setl x \<subseteq> Collect (case_prod R) \<and> setr x \<subseteq> Collect (case_prod S)} (map_sum fst fst))\<inverse>\<inverse> OO |
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
87 |
Grp {x. setl x \<subseteq> Collect (case_prod R) \<and> setr x \<subseteq> Collect (case_prod S)} (map_sum snd snd)" |
58916 | 88 |
unfolding sum_set_defs Grp_def relcompp.simps conversep.simps fun_eq_iff |
89 |
by (fastforce elim: rel_sum.cases split: sum.splits) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
90 |
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
|
91 |
|
58916 | 92 |
inductive_set fsts :: "'a \<times> 'b \<Rightarrow> 'a set" for p :: "'a \<times> 'b" where |
93 |
"fst p \<in> fsts p" |
|
94 |
inductive_set snds :: "'a \<times> 'b \<Rightarrow> 'b set" for p :: "'a \<times> 'b" where |
|
95 |
"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
|
96 |
|
58916 | 97 |
lemma prod_set_defs[code]: "fsts = (\<lambda>p. {fst p})" "snds = (\<lambda>p. {snd p})" |
98 |
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
|
99 |
|
58916 | 100 |
inductive |
101 |
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 | 102 |
where |
58916 | 103 |
"\<lbrakk>R1 a b; R2 c d\<rbrakk> \<Longrightarrow> rel_prod R1 R2 (a, c) (b, d)" |
104 |
||
105 |
lemma rel_prod_apply [code, simp]: |
|
106 |
"rel_prod R1 R2 (a, b) (c, d) \<longleftrightarrow> R1 a c \<and> R2 b d" |
|
107 |
by (auto intro: rel_prod.intros elim: rel_prod.cases) |
|
108 |
||
109 |
lemma rel_prod_conv: |
|
55944 | 110 |
"rel_prod R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)" |
58916 | 111 |
by (rule ext, rule ext) auto |
55083 | 112 |
|
54421 | 113 |
bnf "'a \<times> 'b" |
55932 | 114 |
map: map_prod |
54421 | 115 |
sets: fsts snds |
116 |
bd: natLeq |
|
55944 | 117 |
rel: rel_prod |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
118 |
proof (unfold prod_set_defs) |
55932 | 119 |
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
|
120 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
121 |
fix f1 f2 g1 g2 |
55932 | 122 |
show "map_prod (g1 o f1) (g2 o f2) = map_prod g1 g2 o map_prod f1 f2" |
123 |
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
|
124 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
125 |
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
|
126 |
assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z" |
55932 | 127 |
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
|
128 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
129 |
fix f1 f2 |
55932 | 130 |
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
|
131 |
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
|
132 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
133 |
fix f1 f2 |
55932 | 134 |
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
|
135 |
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
|
136 |
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
|
137 |
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
|
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 |
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
|
140 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
141 |
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
|
142 |
show "|{fst x}| \<le>o natLeq" |
55811 | 143 |
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
|
144 |
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
|
145 |
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
|
146 |
show "|{snd x}| \<le>o natLeq" |
55811 | 147 |
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
|
148 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
149 |
fix R1 R2 S1 S2 |
55944 | 150 |
show "rel_prod R1 R2 OO rel_prod S1 S2 \<le> rel_prod (R1 OO S1) (R2 OO S2)" by auto |
49453 | 151 |
next |
152 |
fix R S |
|
55944 | 153 |
show "rel_prod R S = |
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
154 |
(Grp {x. {fst x} \<subseteq> Collect (case_prod R) \<and> {snd x} \<subseteq> Collect (case_prod S)} (map_prod fst fst))\<inverse>\<inverse> OO |
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
155 |
Grp {x. {fst x} \<subseteq> Collect (case_prod R) \<and> {snd x} \<subseteq> Collect (case_prod S)} (map_prod snd snd)" |
58916 | 156 |
unfolding prod_set_defs rel_prod_apply Grp_def relcompp.simps conversep.simps fun_eq_iff |
49453 | 157 |
by auto |
54189
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
traytel
parents:
53026
diff
changeset
|
158 |
qed |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
159 |
|
54421 | 160 |
bnf "'a \<Rightarrow> 'b" |
161 |
map: "op \<circ>" |
|
162 |
sets: range |
|
163 |
bd: "natLeq +c |UNIV :: 'a set|" |
|
55945 | 164 |
rel: "rel_fun op =" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
165 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
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
|
167 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
168 |
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
|
169 |
unfolding comp_def[abs_def] .. |
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 f g |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
172 |
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
|
173 |
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
|
174 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
175 |
fix f show "range \<circ> op \<circ> f = op ` f \<circ> range" |
56077 | 176 |
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
|
177 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
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
|
179 |
apply (rule card_order_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
180 |
apply (rule natLeq_card_order) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
181 |
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
|
182 |
(* *) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
183 |
show "cinfinite (natLeq +c ?U)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
184 |
apply (rule cinfinite_csum) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
185 |
apply (rule disjI1) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
186 |
by (rule natLeq_cinfinite) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
187 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
188 |
fix f :: "'d => 'a" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
189 |
have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image) |
54486 | 190 |
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
|
191 |
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
|
192 |
next |
54841
af71b753c459
express weak pullback property of bnfs only in terms of the relator
traytel
parents:
54581
diff
changeset
|
193 |
fix R S |
55945 | 194 |
show "rel_fun op = R OO rel_fun op = S \<le> rel_fun op = (R OO S)" by (auto simp: rel_fun_def) |
49453 | 195 |
next |
49463 | 196 |
fix R |
55945 | 197 |
show "rel_fun op = R = |
61032
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
198 |
(Grp {x. range x \<subseteq> Collect (case_prod R)} (op \<circ> fst))\<inverse>\<inverse> OO |
b57df8eecad6
standardized some occurences of ancient "split" alias
haftmann
parents:
60758
diff
changeset
|
199 |
Grp {x. range x \<subseteq> Collect (case_prod R)} (op \<circ> snd)" |
55945 | 200 |
unfolding rel_fun_def Grp_def fun_eq_iff relcompp.simps conversep.simps subset_iff image_iff |
55811 | 201 |
comp_apply mem_Collect_eq split_beta bex_UNIV |
202 |
proof (safe, unfold fun_eq_iff[symmetric]) |
|
203 |
fix x xa a b c xb y aa ba |
|
204 |
assume *: "x = a" "xa = c" "a = ba" "b = aa" "c = (\<lambda>x. snd (b x))" "ba = (\<lambda>x. fst (aa x))" and |
|
205 |
**: "\<forall>t. (\<exists>x. t = aa x) \<longrightarrow> R (fst t) (snd t)" |
|
206 |
show "R (x y) (xa y)" unfolding * by (rule mp[OF spec[OF **]]) blast |
|
207 |
qed force |
|
54189
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
traytel
parents:
53026
diff
changeset
|
208 |
qed |
54191 | 209 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
210 |
end |