| author | nipkow | 
| Fri, 28 Feb 2014 18:09:37 +0100 | |
| changeset 55807 | fd31d0e70eb8 | 
| parent 55707 | 50cf04dd2584 | 
| child 55811 | aa1acc25126b | 
| 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: 
49236diff
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: 
49236diff
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 | |
| 54421 | 16 | bnf ID: 'a | 
| 17 |   map: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | |
| 18 |   sets: "\<lambda>x. {x}"
 | |
| 19 | bd: natLeq | |
| 20 |   rel: "id :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 21 | apply (auto simp: Grp_def fun_eq_iff relcompp.simps natLeq_card_order natLeq_cinfinite) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 22 | apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]]) | 
| 49453 | 23 | apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3] | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 24 | done | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 25 | |
| 54421 | 26 | bnf DEADID: 'a | 
| 27 | map: "id :: 'a \<Rightarrow> 'a" | |
| 55707 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 traytel parents: 
55084diff
changeset | 28 | bd: natLeq | 
| 54421 | 29 | rel: "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 55707 
50cf04dd2584
clarified interaction with dead variables in the composition of BNFs
 traytel parents: 
55084diff
changeset | 30 | by (auto simp add: Grp_def natLeq_card_order natLeq_cinfinite) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 31 | |
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 32 | definition setl :: "'a + 'b \<Rightarrow> 'a set" where | 
| 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 33 | "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 | 34 | |
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 35 | definition setr :: "'a + 'b \<Rightarrow> 'b set" where | 
| 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 36 | "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 | 37 | |
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 38 | 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 | 39 | |
| 55083 | 40 | definition | 
| 41 |    sum_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd \<Rightarrow> bool"
 | |
| 42 | where | |
| 43 | "sum_rel R1 R2 x y = | |
| 44 | (case (x, y) of (Inl x, Inl y) \<Rightarrow> R1 x y | |
| 45 | | (Inr x, Inr y) \<Rightarrow> R2 x y | |
| 46 | | _ \<Rightarrow> False)" | |
| 47 | ||
| 48 | lemma sum_rel_simps[simp]: | |
| 49 | "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1" | |
| 50 | "sum_rel R1 R2 (Inl a1) (Inr b2) = False" | |
| 51 | "sum_rel R1 R2 (Inr a2) (Inl b1) = False" | |
| 52 | "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" | |
| 53 | unfolding sum_rel_def by simp_all | |
| 54 | ||
| 54421 | 55 | bnf "'a + 'b" | 
| 56 | map: sum_map | |
| 57 | sets: setl setr | |
| 58 | bd: natLeq | |
| 59 | wits: Inl Inr | |
| 60 | rel: sum_rel | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 61 | proof - | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 62 | show "sum_map id id = id" by (rule sum_map.id) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 63 | next | 
| 54486 | 64 | fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r" | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 65 | show "sum_map (g1 o f1) (g2 o f2) = sum_map g1 g2 o sum_map f1 f2" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 66 | by (rule sum_map.comp[symmetric]) | 
| 
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 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: 
49450diff
changeset | 69 | 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: 
49450diff
changeset | 70 | a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z" | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 71 | thus "sum_map f1 f2 x = sum_map g1 g2 x" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 72 | proof (cases x) | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 73 | 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 | 74 | next | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 75 | 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 | 76 | qed | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 77 | next | 
| 54486 | 78 | fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 79 | show "setl o sum_map f1 f2 = image f1 o setl" | 
| 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 80 | 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 | 81 | next | 
| 54486 | 82 | fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 83 | show "setr o sum_map f1 f2 = image f2 o setr" | 
| 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 84 | 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 | 85 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 86 | 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 | 87 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 88 | 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 | 89 | next | 
| 54486 | 90 | fix x :: "'o + 'p" | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 91 | 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 | 92 | apply (rule ordLess_imp_ordLeq) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 93 | apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 94 | 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 | 95 | next | 
| 54486 | 96 | fix x :: "'o + 'p" | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 97 | 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 | 98 | apply (rule ordLess_imp_ordLeq) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 99 | apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) | 
| 49451 
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
 blanchet parents: 
49450diff
changeset | 100 | 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 | 101 | next | 
| 54841 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 102 | fix R1 R2 S1 S2 | 
| 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 103 | show "sum_rel R1 R2 OO sum_rel S1 S2 \<le> sum_rel (R1 OO S1) (R2 OO S2)" | 
| 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 104 | by (auto simp: sum_rel_def split: sum.splits) | 
| 49453 | 105 | next | 
| 106 | fix R S | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 107 | show "sum_rel R S = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 108 |         (Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (sum_map fst fst))\<inverse>\<inverse> OO
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 109 |         Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (sum_map snd snd)"
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 110 | unfolding setl_def setr_def sum_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff | 
| 49453 | 111 | 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 | 112 | 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 | 113 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 114 | 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 | 115 | "fsts x = {fst x}"
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 116 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 117 | 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 | 118 | "snds x = {snd x}"
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 119 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 120 | 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 | 121 | |
| 55083 | 122 | definition | 
| 123 |   prod_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b \<times> 'd \<Rightarrow> bool"
 | |
| 124 | where | |
| 125 | "prod_rel R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)" | |
| 126 | ||
| 127 | lemma prod_rel_apply [simp]: | |
| 128 | "prod_rel R1 R2 (a, b) (c, d) \<longleftrightarrow> R1 a c \<and> R2 b d" | |
| 129 | by (simp add: prod_rel_def) | |
| 130 | ||
| 54421 | 131 | bnf "'a \<times> 'b" | 
| 132 | map: map_pair | |
| 133 | sets: fsts snds | |
| 134 | bd: natLeq | |
| 135 | rel: prod_rel | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 136 | proof (unfold prod_set_defs) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 137 | show "map_pair id id = id" by (rule map_pair.id) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 138 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 139 | fix f1 f2 g1 g2 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 140 | show "map_pair (g1 o f1) (g2 o f2) = map_pair g1 g2 o map_pair f1 f2" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 141 | by (rule map_pair.comp[symmetric]) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 142 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 143 | 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 | 144 |   assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z"
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 145 | thus "map_pair f1 f2 x = map_pair g1 g2 x" by (cases x) simp | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 146 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 147 | fix f1 f2 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 148 |   show "(\<lambda>x. {fst x}) o map_pair f1 f2 = image f1 o (\<lambda>x. {fst x})"
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 149 | 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 | 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 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 152 |   show "(\<lambda>x. {snd x}) o map_pair f1 f2 = image f2 o (\<lambda>x. {snd x})"
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 153 | 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 | 154 | next | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 155 | 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 | 156 | next | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 157 | 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 | 158 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 159 | fix x | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 160 |   show "|{fst x}| \<le>o natLeq"
 | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 161 | by (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq finite.emptyI finite_insert) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 162 | next | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 163 | fix x | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 164 |   show "|{snd x}| \<le>o natLeq"
 | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
52545diff
changeset | 165 | by (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq finite.emptyI finite_insert) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 166 | next | 
| 54841 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 167 | fix R1 R2 S1 S2 | 
| 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 168 | show "prod_rel R1 R2 OO prod_rel S1 S2 \<le> prod_rel (R1 OO S1) (R2 OO S2)" by auto | 
| 49453 | 169 | next | 
| 170 | fix R S | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 171 | show "prod_rel R S = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 172 |         (Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair fst fst))\<inverse>\<inverse> OO
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 173 |         Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair snd snd)"
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 174 | unfolding prod_set_defs prod_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff | 
| 49453 | 175 | by auto | 
| 54189 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 traytel parents: 
53026diff
changeset | 176 | qed | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 177 | |
| 54421 | 178 | bnf "'a \<Rightarrow> 'b" | 
| 179 | map: "op \<circ>" | |
| 180 | sets: range | |
| 181 | bd: "natLeq +c |UNIV :: 'a set|" | |
| 182 | rel: "fun_rel op =" | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 183 | proof | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 184 | 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 | 185 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 186 | 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 | 187 | unfolding comp_def[abs_def] .. | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 188 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 189 | fix x f g | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 190 | 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 | 191 | 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 | 192 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 193 | fix f show "range \<circ> op \<circ> f = op ` f \<circ> range" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 194 | unfolding image_def comp_def[abs_def] by auto | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 195 | next | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 196 | 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 | 197 | apply (rule card_order_csum) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 198 | apply (rule natLeq_card_order) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 199 | 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 | 200 | (* *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 201 | show "cinfinite (natLeq +c ?U)" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 202 | apply (rule cinfinite_csum) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 203 | apply (rule disjI1) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 204 | by (rule natLeq_cinfinite) | 
| 
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 :: "'d => 'a" | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 207 | have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image) | 
| 54486 | 208 | 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 | 209 | 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 | 210 | next | 
| 54841 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 211 | fix R S | 
| 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 traytel parents: 
54581diff
changeset | 212 | show "fun_rel op = R OO fun_rel op = S \<le> fun_rel op = (R OO S)" by (auto simp: fun_rel_def) | 
| 49453 | 213 | next | 
| 49463 | 214 | fix R | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 215 | show "fun_rel op = R = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 216 |         (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: 
51836diff
changeset | 217 |          Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)"
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 218 | unfolding fun_rel_def Grp_def fun_eq_iff relcompp.simps conversep.simps subset_iff image_iff | 
| 54486 | 219 | by auto (force, metis (no_types) pair_collapse) | 
| 54189 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 traytel parents: 
53026diff
changeset | 220 | qed | 
| 54191 | 221 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 222 | end |