author | immler |
Tue, 15 May 2018 11:33:43 +0200 | |
changeset 68188 | 2af1f142f855 |
parent 68028 | 1f9f973eed2a |
child 68406 | 6beb45f6cf67 |
permissions | -rw-r--r-- |
37653 | 1 |
(* Title: HOL/Library/Cardinality.thy |
48051 | 2 |
Author: Brian Huffman, Andreas Lochbihler |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
3 |
*) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
4 |
|
60500 | 5 |
section \<open>Cardinality of types\<close> |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
6 |
|
37653 | 7 |
theory Cardinality |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
8 |
imports Phantom_Type |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
9 |
begin |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
10 |
|
60500 | 11 |
subsection \<open>Preliminary lemmas\<close> |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
12 |
(* These should be moved elsewhere *) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
13 |
|
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
14 |
lemma (in type_definition) univ: |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
15 |
"UNIV = Abs ` A" |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
16 |
proof |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
17 |
show "Abs ` A \<subseteq> UNIV" by (rule subset_UNIV) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
18 |
show "UNIV \<subseteq> Abs ` A" |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
19 |
proof |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
20 |
fix x :: 'b |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
21 |
have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
22 |
moreover have "Rep x \<in> A" by (rule Rep) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
23 |
ultimately show "x \<in> Abs ` A" by (rule image_eqI) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
24 |
qed |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
25 |
qed |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
26 |
|
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
27 |
lemma (in type_definition) card: "card (UNIV :: 'b set) = card A" |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
28 |
by (simp add: univ card_image inj_on_def Abs_inject) |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
29 |
|
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
30 |
|
60500 | 31 |
subsection \<open>Cardinalities of types\<close> |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
32 |
|
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
33 |
syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))") |
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
34 |
|
61076 | 35 |
translations "CARD('t)" => "CONST card (CONST UNIV :: 't set)" |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
36 |
|
60500 | 37 |
print_translation \<open> |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
38 |
let |
52147 | 39 |
fun card_univ_tr' ctxt [Const (@{const_syntax UNIV}, Type (_, [T]))] = |
49689
b8a710806de9
recovered print translation after 'a => bool to 'a set change;
wenzelm
parents:
48176
diff
changeset
|
40 |
Syntax.const @{syntax_const "_type_card"} $ Syntax_Phases.term_of_typ ctxt T |
42247
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
wenzelm
parents:
42245
diff
changeset
|
41 |
in [(@{const_syntax card}, card_univ_tr')] end |
60500 | 42 |
\<close> |
24407 | 43 |
|
48058 | 44 |
lemma card_prod [simp]: "CARD('a \<times> 'b) = CARD('a) * CARD('b)" |
26153 | 45 |
unfolding UNIV_Times_UNIV [symmetric] by (simp only: card_cartesian_product) |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
46 |
|
48060 | 47 |
lemma card_UNIV_sum: "CARD('a + 'b) = (if CARD('a) \<noteq> 0 \<and> CARD('b) \<noteq> 0 then CARD('a) + CARD('b) else 0)" |
48 |
unfolding UNIV_Plus_UNIV[symmetric] |
|
49 |
by(auto simp add: card_eq_0_iff card_Plus simp del: UNIV_Plus_UNIV) |
|
50 |
||
30001 | 51 |
lemma card_sum [simp]: "CARD('a + 'b) = CARD('a::finite) + CARD('b::finite)" |
48060 | 52 |
by(simp add: card_UNIV_sum) |
53 |
||
54 |
lemma card_UNIV_option: "CARD('a option) = (if CARD('a) = 0 then 0 else CARD('a) + 1)" |
|
55 |
proof - |
|
56 |
have "(None :: 'a option) \<notin> range Some" by clarsimp |
|
57 |
thus ?thesis |
|
53191 | 58 |
by (simp add: UNIV_option_conv card_eq_0_iff finite_range_Some card_image) |
48060 | 59 |
qed |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
60 |
|
30001 | 61 |
lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)" |
48060 | 62 |
by(simp add: card_UNIV_option) |
63 |
||
64 |
lemma card_UNIV_set: "CARD('a set) = (if CARD('a) = 0 then 0 else 2 ^ CARD('a))" |
|
65 |
by(simp add: Pow_UNIV[symmetric] card_eq_0_iff card_Pow del: Pow_UNIV) |
|
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
66 |
|
30001 | 67 |
lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)" |
48060 | 68 |
by(simp add: card_UNIV_set) |
24332
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff
changeset
|
69 |
|
30001 | 70 |
lemma card_nat [simp]: "CARD(nat) = 0" |
44142 | 71 |
by (simp add: card_eq_0_iff) |
30001 | 72 |
|
48060 | 73 |
lemma card_fun: "CARD('a \<Rightarrow> 'b) = (if CARD('a) \<noteq> 0 \<and> CARD('b) \<noteq> 0 \<or> CARD('b) = 1 then CARD('b) ^ CARD('a) else 0)" |
74 |
proof - |
|
75 |
{ assume "0 < CARD('a)" and "0 < CARD('b)" |
|
76 |
hence fina: "finite (UNIV :: 'a set)" and finb: "finite (UNIV :: 'b set)" |
|
77 |
by(simp_all only: card_ge_0_finite) |
|
78 |
from finite_distinct_list[OF finb] obtain bs |
|
79 |
where bs: "set bs = (UNIV :: 'b set)" and distb: "distinct bs" by blast |
|
80 |
from finite_distinct_list[OF fina] obtain as |
|
81 |
where as: "set as = (UNIV :: 'a set)" and dista: "distinct as" by blast |
|
82 |
have cb: "CARD('b) = length bs" |
|
83 |
unfolding bs[symmetric] distinct_card[OF distb] .. |
|
84 |
have ca: "CARD('a) = length as" |
|
85 |
unfolding as[symmetric] distinct_card[OF dista] .. |
|
67091 | 86 |
let ?xs = "map (\<lambda>ys. the \<circ> map_of (zip as ys)) (List.n_lists (length as) bs)" |
48060 | 87 |
have "UNIV = set ?xs" |
88 |
proof(rule UNIV_eq_I) |
|
89 |
fix f :: "'a \<Rightarrow> 'b" |
|
90 |
from as have "f = the \<circ> map_of (zip as (map f as))" |
|
91 |
by(auto simp add: map_of_zip_map) |
|
92 |
thus "f \<in> set ?xs" using bs by(auto simp add: set_n_lists) |
|
93 |
qed |
|
94 |
moreover have "distinct ?xs" unfolding distinct_map |
|
95 |
proof(intro conjI distinct_n_lists distb inj_onI) |
|
96 |
fix xs ys :: "'b list" |
|
49948
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49689
diff
changeset
|
97 |
assume xs: "xs \<in> set (List.n_lists (length as) bs)" |
744934b818c7
moved quite generic material from theory Enum to more appropriate places
haftmann
parents:
49689
diff
changeset
|
98 |
and ys: "ys \<in> set (List.n_lists (length as) bs)" |
48060 | 99 |
and eq: "the \<circ> map_of (zip as xs) = the \<circ> map_of (zip as ys)" |
100 |
from xs ys have [simp]: "length xs = length as" "length ys = length as" |
|
101 |
by(simp_all add: length_n_lists_elem) |
|
102 |
have "map_of (zip as xs) = map_of (zip as ys)" |
|
103 |
proof |
|
104 |
fix x |
|
105 |
from as bs have "\<exists>y. map_of (zip as xs) x = Some y" "\<exists>y. map_of (zip as ys) x = Some y" |
|
106 |
by(simp_all add: map_of_zip_is_Some[symmetric]) |
|
107 |
with eq show "map_of (zip as xs) x = map_of (zip as ys) x" |
|
108 |
by(auto dest: fun_cong[where x=x]) |
|
109 |
qed |
|
110 |
with dista show "xs = ys" by(simp add: map_of_zip_inject) |
|
111 |
qed |
|
112 |
hence "card (set ?xs) = length ?xs" by(simp only: distinct_card) |
|
113 |
moreover have "length ?xs = length bs ^ length as" by(simp add: length_n_lists) |
|
114 |
ultimately have "CARD('a \<Rightarrow> 'b) = CARD('b) ^ CARD('a)" using cb ca by simp } |
|
115 |
moreover { |
|
116 |
assume cb: "CARD('b) = 1" |
|
117 |
then obtain b where b: "UNIV = {b :: 'b}" by(auto simp add: card_Suc_eq) |
|
118 |
have eq: "UNIV = {\<lambda>x :: 'a. b ::'b}" |
|
119 |
proof(rule UNIV_eq_I) |
|
120 |
fix x :: "'a \<Rightarrow> 'b" |
|
121 |
{ fix y |
|
122 |
have "x y \<in> UNIV" .. |
|
123 |
hence "x y = b" unfolding b by simp } |
|
124 |
thus "x \<in> {\<lambda>x. b}" by(auto) |
|
125 |
qed |
|
126 |
have "CARD('a \<Rightarrow> 'b) = 1" unfolding eq by simp } |
|
127 |
ultimately show ?thesis |
|
128 |
by(auto simp del: One_nat_def)(auto simp add: card_eq_0_iff dest: finite_fun_UNIVD2 finite_fun_UNIVD1) |
|
129 |
qed |
|
130 |
||
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
131 |
corollary finite_UNIV_fun: |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
132 |
"finite (UNIV :: ('a \<Rightarrow> 'b) set) \<longleftrightarrow> |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
133 |
finite (UNIV :: 'a set) \<and> finite (UNIV :: 'b set) \<or> CARD('b) = 1" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
134 |
(is "?lhs \<longleftrightarrow> ?rhs") |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
135 |
proof - |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
136 |
have "?lhs \<longleftrightarrow> CARD('a \<Rightarrow> 'b) > 0" by(simp add: card_gt_0_iff) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
137 |
also have "\<dots> \<longleftrightarrow> CARD('a) > 0 \<and> CARD('b) > 0 \<or> CARD('b) = 1" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
138 |
by(simp add: card_fun) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
139 |
also have "\<dots> = ?rhs" by(simp add: card_gt_0_iff) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
140 |
finally show ?thesis . |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
141 |
qed |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
142 |
|
48060 | 143 |
lemma card_literal: "CARD(String.literal) = 0" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
144 |
by(simp add: card_eq_0_iff infinite_literal) |
30001 | 145 |
|
60500 | 146 |
subsection \<open>Classes with at least 1 and 2\<close> |
30001 | 147 |
|
60500 | 148 |
text \<open>Class finite already captures "at least 1"\<close> |
30001 | 149 |
|
150 |
lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)" |
|
29997 | 151 |
unfolding neq0_conv [symmetric] by simp |
152 |
||
30001 | 153 |
lemma one_le_card_finite [simp]: "Suc 0 \<le> CARD('a::finite)" |
154 |
by (simp add: less_Suc_eq_le [symmetric]) |
|
155 |
||
60500 | 156 |
text \<open>Class for cardinality "at least 2"\<close> |
30001 | 157 |
|
158 |
class card2 = finite + |
|
159 |
assumes two_le_card: "2 \<le> CARD('a)" |
|
160 |
||
161 |
lemma one_less_card: "Suc 0 < CARD('a::card2)" |
|
162 |
using two_le_card [where 'a='a] by simp |
|
163 |
||
164 |
lemma one_less_int_card: "1 < int CARD('a::card2)" |
|
165 |
using one_less_card [where 'a='a] by simp |
|
166 |
||
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
167 |
|
60500 | 168 |
subsection \<open>A type class for deciding finiteness of types\<close> |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
169 |
|
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
170 |
type_synonym 'a finite_UNIV = "('a, bool) phantom" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
171 |
|
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
172 |
class finite_UNIV = |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
173 |
fixes finite_UNIV :: "('a, bool) phantom" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
174 |
assumes finite_UNIV: "finite_UNIV = Phantom('a) (finite (UNIV :: 'a set))" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
175 |
|
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
176 |
lemma finite_UNIV_code [code_unfold]: |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
177 |
"finite (UNIV :: 'a :: finite_UNIV set) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
178 |
\<longleftrightarrow> of_phantom (finite_UNIV :: 'a finite_UNIV)" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
179 |
by(simp add: finite_UNIV) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
180 |
|
60500 | 181 |
subsection \<open>A type class for computing the cardinality of types\<close> |
48051 | 182 |
|
48059 | 183 |
definition is_list_UNIV :: "'a list \<Rightarrow> bool" |
51116
0dac0158b8d4
implement code generation for finite, card, op = and op <= for sets always via finite_UNIV and card_UNIV, as fragile rewrites based on sorts are hard to find and debug
Andreas Lochbihler
parents:
49948
diff
changeset
|
184 |
where "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)" |
48059 | 185 |
|
186 |
lemma is_list_UNIV_iff: "is_list_UNIV xs \<longleftrightarrow> set xs = UNIV" |
|
187 |
by(auto simp add: is_list_UNIV_def Let_def card_eq_0_iff List.card_set[symmetric] |
|
188 |
dest: subst[where P="finite", OF _ finite_set] card_eq_UNIV_imp_eq_UNIV) |
|
189 |
||
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
190 |
type_synonym 'a card_UNIV = "('a, nat) phantom" |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
191 |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
192 |
class card_UNIV = finite_UNIV + |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
193 |
fixes card_UNIV :: "'a card_UNIV" |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
194 |
assumes card_UNIV: "card_UNIV = Phantom('a) CARD('a)" |
48051 | 195 |
|
61585 | 196 |
subsection \<open>Instantiations for \<open>card_UNIV\<close>\<close> |
48051 | 197 |
|
198 |
instantiation nat :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
199 |
definition "finite_UNIV = Phantom(nat) False" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
200 |
definition "card_UNIV = Phantom(nat) 0" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
201 |
instance by intro_classes (simp_all add: finite_UNIV_nat_def card_UNIV_nat_def) |
48051 | 202 |
end |
203 |
||
204 |
instantiation int :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
205 |
definition "finite_UNIV = Phantom(int) False" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
206 |
definition "card_UNIV = Phantom(int) 0" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
207 |
instance by intro_classes (simp_all add: card_UNIV_int_def finite_UNIV_int_def infinite_UNIV_int) |
48051 | 208 |
end |
209 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
210 |
instantiation natural :: card_UNIV begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
211 |
definition "finite_UNIV = Phantom(natural) False" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
212 |
definition "card_UNIV = Phantom(natural) 0" |
60679 | 213 |
instance |
214 |
by standard |
|
215 |
(auto simp add: finite_UNIV_natural_def card_UNIV_natural_def card_eq_0_iff |
|
216 |
type_definition.univ [OF type_definition_natural] natural_eq_iff |
|
217 |
dest!: finite_imageD intro: inj_onI) |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
218 |
end |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
219 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
220 |
instantiation integer :: card_UNIV begin |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
221 |
definition "finite_UNIV = Phantom(integer) False" |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51139
diff
changeset
|
222 |
definition "card_UNIV = Phantom(integer) 0" |
60679 | 223 |
instance |
224 |
by standard |
|
225 |
(auto simp add: finite_UNIV_integer_def card_UNIV_integer_def card_eq_0_iff |
|
226 |
type_definition.univ [OF type_definition_integer] infinite_UNIV_int |
|
227 |
dest!: finite_imageD intro: inj_onI) |
|
48165
d07a0b9601aa
instantiate card_UNIV with nibble and code_numeral
Andreas Lochbihler
parents:
48164
diff
changeset
|
228 |
end |
d07a0b9601aa
instantiate card_UNIV with nibble and code_numeral
Andreas Lochbihler
parents:
48164
diff
changeset
|
229 |
|
48051 | 230 |
instantiation list :: (type) card_UNIV begin |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
231 |
definition "finite_UNIV = Phantom('a list) False" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
232 |
definition "card_UNIV = Phantom('a list) 0" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
233 |
instance by intro_classes (simp_all add: card_UNIV_list_def finite_UNIV_list_def infinite_UNIV_listI) |
48051 | 234 |
end |
235 |
||
236 |
instantiation unit :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
237 |
definition "finite_UNIV = Phantom(unit) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
238 |
definition "card_UNIV = Phantom(unit) 1" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
239 |
instance by intro_classes (simp_all add: card_UNIV_unit_def finite_UNIV_unit_def) |
48051 | 240 |
end |
241 |
||
242 |
instantiation bool :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
243 |
definition "finite_UNIV = Phantom(bool) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
244 |
definition "card_UNIV = Phantom(bool) 2" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
245 |
instance by(intro_classes)(simp_all add: card_UNIV_bool_def finite_UNIV_bool_def) |
48051 | 246 |
end |
247 |
||
248 |
instantiation char :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
249 |
definition "finite_UNIV = Phantom(char) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
250 |
definition "card_UNIV = Phantom(char) 256" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
251 |
instance by intro_classes (simp_all add: card_UNIV_char_def card_UNIV_char finite_UNIV_char_def) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
252 |
end |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
253 |
|
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
254 |
instantiation prod :: (finite_UNIV, finite_UNIV) finite_UNIV begin |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
255 |
definition "finite_UNIV = Phantom('a \<times> 'b) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
256 |
(of_phantom (finite_UNIV :: 'a finite_UNIV) \<and> of_phantom (finite_UNIV :: 'b finite_UNIV))" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
257 |
instance by intro_classes (simp add: finite_UNIV_prod_def finite_UNIV finite_prod) |
48051 | 258 |
end |
259 |
||
260 |
instantiation prod :: (card_UNIV, card_UNIV) card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
261 |
definition "card_UNIV = Phantom('a \<times> 'b) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
262 |
(of_phantom (card_UNIV :: 'a card_UNIV) * of_phantom (card_UNIV :: 'b card_UNIV))" |
48060 | 263 |
instance by intro_classes (simp add: card_UNIV_prod_def card_UNIV) |
48051 | 264 |
end |
265 |
||
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
266 |
instantiation sum :: (finite_UNIV, finite_UNIV) finite_UNIV begin |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
267 |
definition "finite_UNIV = Phantom('a + 'b) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
268 |
(of_phantom (finite_UNIV :: 'a finite_UNIV) \<and> of_phantom (finite_UNIV :: 'b finite_UNIV))" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
269 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
270 |
by intro_classes (simp add: UNIV_Plus_UNIV[symmetric] finite_UNIV_sum_def finite_UNIV del: UNIV_Plus_UNIV) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
271 |
end |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
272 |
|
48051 | 273 |
instantiation sum :: (card_UNIV, card_UNIV) card_UNIV begin |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
274 |
definition "card_UNIV = Phantom('a + 'b) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
275 |
(let ca = of_phantom (card_UNIV :: 'a card_UNIV); |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
276 |
cb = of_phantom (card_UNIV :: 'b card_UNIV) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
277 |
in if ca \<noteq> 0 \<and> cb \<noteq> 0 then ca + cb else 0)" |
48060 | 278 |
instance by intro_classes (auto simp add: card_UNIV_sum_def card_UNIV card_UNIV_sum) |
48051 | 279 |
end |
280 |
||
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
281 |
instantiation "fun" :: (finite_UNIV, card_UNIV) finite_UNIV begin |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
282 |
definition "finite_UNIV = Phantom('a \<Rightarrow> 'b) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
283 |
(let cb = of_phantom (card_UNIV :: 'b card_UNIV) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
284 |
in cb = 1 \<or> of_phantom (finite_UNIV :: 'a finite_UNIV) \<and> cb \<noteq> 0)" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
285 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
286 |
by intro_classes (auto simp add: finite_UNIV_fun_def Let_def card_UNIV finite_UNIV finite_UNIV_fun card_gt_0_iff) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
287 |
end |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
288 |
|
48051 | 289 |
instantiation "fun" :: (card_UNIV, card_UNIV) card_UNIV begin |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
290 |
definition "card_UNIV = Phantom('a \<Rightarrow> 'b) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
291 |
(let ca = of_phantom (card_UNIV :: 'a card_UNIV); |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
292 |
cb = of_phantom (card_UNIV :: 'b card_UNIV) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
293 |
in if ca \<noteq> 0 \<and> cb \<noteq> 0 \<or> cb = 1 then cb ^ ca else 0)" |
48060 | 294 |
instance by intro_classes (simp add: card_UNIV_fun_def card_UNIV Let_def card_fun) |
295 |
end |
|
48051 | 296 |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
297 |
instantiation option :: (finite_UNIV) finite_UNIV begin |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
298 |
definition "finite_UNIV = Phantom('a option) (of_phantom (finite_UNIV :: 'a finite_UNIV))" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
299 |
instance by intro_classes (simp add: finite_UNIV_option_def finite_UNIV) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
300 |
end |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
301 |
|
48060 | 302 |
instantiation option :: (card_UNIV) card_UNIV begin |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
303 |
definition "card_UNIV = Phantom('a option) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
304 |
(let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c \<noteq> 0 then Suc c else 0)" |
48060 | 305 |
instance by intro_classes (simp add: card_UNIV_option_def card_UNIV card_UNIV_option) |
306 |
end |
|
48051 | 307 |
|
48060 | 308 |
instantiation String.literal :: card_UNIV begin |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
309 |
definition "finite_UNIV = Phantom(String.literal) False" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
310 |
definition "card_UNIV = Phantom(String.literal) 0" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
311 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
312 |
by intro_classes (simp_all add: card_UNIV_literal_def finite_UNIV_literal_def infinite_literal card_literal) |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
313 |
end |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
314 |
|
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
315 |
instantiation set :: (finite_UNIV) finite_UNIV begin |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
316 |
definition "finite_UNIV = Phantom('a set) (of_phantom (finite_UNIV :: 'a finite_UNIV))" |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
317 |
instance by intro_classes (simp add: finite_UNIV_set_def finite_UNIV Finite_Set.finite_set) |
48060 | 318 |
end |
319 |
||
320 |
instantiation set :: (card_UNIV) card_UNIV begin |
|
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
321 |
definition "card_UNIV = Phantom('a set) |
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
322 |
(let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c = 0 then 0 else 2 ^ c)" |
48060 | 323 |
instance by intro_classes (simp add: card_UNIV_set_def card_UNIV_set card_UNIV) |
48051 | 324 |
end |
325 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52147
diff
changeset
|
326 |
lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^sub>1]" |
48060 | 327 |
by(auto intro: finite_1.exhaust) |
328 |
||
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52147
diff
changeset
|
329 |
lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^sub>1, finite_2.a\<^sub>2]" |
48060 | 330 |
by(auto intro: finite_2.exhaust) |
48051 | 331 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52147
diff
changeset
|
332 |
lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^sub>1, finite_3.a\<^sub>2, finite_3.a\<^sub>3]" |
48060 | 333 |
by(auto intro: finite_3.exhaust) |
48051 | 334 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52147
diff
changeset
|
335 |
lemma UNIV_finite_4: "UNIV = set [finite_4.a\<^sub>1, finite_4.a\<^sub>2, finite_4.a\<^sub>3, finite_4.a\<^sub>4]" |
48060 | 336 |
by(auto intro: finite_4.exhaust) |
337 |
||
338 |
lemma UNIV_finite_5: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52147
diff
changeset
|
339 |
"UNIV = set [finite_5.a\<^sub>1, finite_5.a\<^sub>2, finite_5.a\<^sub>3, finite_5.a\<^sub>4, finite_5.a\<^sub>5]" |
48060 | 340 |
by(auto intro: finite_5.exhaust) |
48051 | 341 |
|
48060 | 342 |
instantiation Enum.finite_1 :: card_UNIV begin |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
343 |
definition "finite_UNIV = Phantom(Enum.finite_1) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
344 |
definition "card_UNIV = Phantom(Enum.finite_1) 1" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
345 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
346 |
by intro_classes (simp_all add: UNIV_finite_1 card_UNIV_finite_1_def finite_UNIV_finite_1_def) |
48060 | 347 |
end |
348 |
||
349 |
instantiation Enum.finite_2 :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
350 |
definition "finite_UNIV = Phantom(Enum.finite_2) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
351 |
definition "card_UNIV = Phantom(Enum.finite_2) 2" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
352 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
353 |
by intro_classes (simp_all add: UNIV_finite_2 card_UNIV_finite_2_def finite_UNIV_finite_2_def) |
48060 | 354 |
end |
48051 | 355 |
|
48060 | 356 |
instantiation Enum.finite_3 :: card_UNIV begin |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
357 |
definition "finite_UNIV = Phantom(Enum.finite_3) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
358 |
definition "card_UNIV = Phantom(Enum.finite_3) 3" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
359 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
360 |
by intro_classes (simp_all add: UNIV_finite_3 card_UNIV_finite_3_def finite_UNIV_finite_3_def) |
48060 | 361 |
end |
362 |
||
363 |
instantiation Enum.finite_4 :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
364 |
definition "finite_UNIV = Phantom(Enum.finite_4) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
365 |
definition "card_UNIV = Phantom(Enum.finite_4) 4" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
366 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
367 |
by intro_classes (simp_all add: UNIV_finite_4 card_UNIV_finite_4_def finite_UNIV_finite_4_def) |
48060 | 368 |
end |
369 |
||
370 |
instantiation Enum.finite_5 :: card_UNIV begin |
|
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
371 |
definition "finite_UNIV = Phantom(Enum.finite_5) True" |
48164
e97369f20c30
change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents:
48070
diff
changeset
|
372 |
definition "card_UNIV = Phantom(Enum.finite_5) 5" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
373 |
instance |
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
374 |
by intro_classes (simp_all add: UNIV_finite_5 card_UNIV_finite_5_def finite_UNIV_finite_5_def) |
48051 | 375 |
end |
376 |
||
60500 | 377 |
subsection \<open>Code setup for sets\<close> |
48051 | 378 |
|
60500 | 379 |
text \<open> |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
380 |
Implement @{term "CARD('a)"} via @{term card_UNIV} and provide |
67399 | 381 |
implementations for @{term "finite"}, @{term "card"}, @{term "(\<subseteq>)"}, |
382 |
and @{term "(=)"}if the calling context already provides @{class finite_UNIV} |
|
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
383 |
and @{class card_UNIV} instances. If we implemented the latter |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
384 |
always via @{term card_UNIV}, we would require instances of essentially all |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
385 |
element types, i.e., a lot of instantiation proofs and -- at run time -- |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
386 |
possibly slow dictionary constructions. |
60500 | 387 |
\<close> |
51116
0dac0158b8d4
implement code generation for finite, card, op = and op <= for sets always via finite_UNIV and card_UNIV, as fragile rewrites based on sorts are hard to find and debug
Andreas Lochbihler
parents:
49948
diff
changeset
|
388 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
389 |
context |
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
390 |
begin |
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
391 |
|
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
392 |
qualified definition card_UNIV' :: "'a card_UNIV" |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
393 |
where [code del]: "card_UNIV' = Phantom('a) CARD('a)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
394 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
395 |
lemma CARD_code [code_unfold]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
396 |
"CARD('a) = of_phantom (card_UNIV' :: 'a card_UNIV)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
397 |
by(simp add: card_UNIV'_def) |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
398 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
399 |
lemma card_UNIV'_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
400 |
"card_UNIV' = card_UNIV" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
401 |
by(simp add: card_UNIV card_UNIV'_def) |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
402 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
403 |
end |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
404 |
|
48051 | 405 |
lemma card_Compl: |
406 |
"finite A \<Longrightarrow> card (- A) = card (UNIV :: 'a set) - card (A :: 'a set)" |
|
407 |
by (metis Compl_eq_Diff_UNIV card_Diff_subset top_greatest) |
|
408 |
||
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
409 |
context fixes xs :: "'a :: finite_UNIV list" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
410 |
begin |
48062
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
411 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
412 |
qualified definition finite' :: "'a set \<Rightarrow> bool" |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
413 |
where [simp, code del, code_abbrev]: "finite' = finite" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
414 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
415 |
lemma finite'_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
416 |
"finite' (set xs) \<longleftrightarrow> True" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
417 |
"finite' (List.coset xs) \<longleftrightarrow> of_phantom (finite_UNIV :: 'a finite_UNIV)" |
48176
3d9c1ddb9f47
new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents:
48165
diff
changeset
|
418 |
by(simp_all add: card_gt_0_iff finite_UNIV) |
48062
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
419 |
|
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
420 |
end |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
421 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
422 |
context fixes xs :: "'a :: card_UNIV list" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
423 |
begin |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
424 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
425 |
qualified definition card' :: "'a set \<Rightarrow> nat" |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
426 |
where [simp, code del, code_abbrev]: "card' = card" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
427 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
428 |
lemma card'_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
429 |
"card' (set xs) = length (remdups xs)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
430 |
"card' (List.coset xs) = of_phantom (card_UNIV :: 'a card_UNIV) - length (remdups xs)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
431 |
by(simp_all add: List.card_set card_Compl card_UNIV) |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
432 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
433 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
434 |
qualified definition subset' :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" |
67399 | 435 |
where [simp, code del, code_abbrev]: "subset' = (\<subseteq>)" |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
436 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
437 |
lemma subset'_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
438 |
"subset' A (List.coset ys) \<longleftrightarrow> (\<forall>y \<in> set ys. y \<notin> A)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
439 |
"subset' (set ys) B \<longleftrightarrow> (\<forall>y \<in> set ys. y \<in> B)" |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
440 |
"subset' (List.coset xs) (set ys) \<longleftrightarrow> (let n = CARD('a) in n > 0 \<and> card(set (xs @ ys)) = n)" |
48062
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
441 |
by(auto simp add: Let_def card_gt_0_iff dest: card_eq_UNIV_imp_eq_UNIV intro: arg_cong[where f=card]) |
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
442 |
(metis finite_compl finite_set rev_finite_subset) |
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
443 |
|
61115
3a4400985780
modernized name space management -- more uniform qualification;
wenzelm
parents:
61076
diff
changeset
|
444 |
qualified definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" |
67399 | 445 |
where [simp, code del, code_abbrev]: "eq_set = (=)" |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
446 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
447 |
lemma eq_set_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
448 |
fixes ys |
48051 | 449 |
defines "rhs \<equiv> |
48059 | 450 |
let n = CARD('a) |
48051 | 451 |
in if n = 0 then False else |
452 |
let xs' = remdups xs; ys' = remdups ys |
|
453 |
in length xs' + length ys' = n \<and> (\<forall>x \<in> set xs'. x \<notin> set ys') \<and> (\<forall>y \<in> set ys'. y \<notin> set xs')" |
|
60583 | 454 |
shows "eq_set (List.coset xs) (set ys) \<longleftrightarrow> rhs" |
455 |
and "eq_set (set ys) (List.coset xs) \<longleftrightarrow> rhs" |
|
456 |
and "eq_set (set xs) (set ys) \<longleftrightarrow> (\<forall>x \<in> set xs. x \<in> set ys) \<and> (\<forall>y \<in> set ys. y \<in> set xs)" |
|
457 |
and "eq_set (List.coset xs) (List.coset ys) \<longleftrightarrow> (\<forall>x \<in> set xs. x \<in> set ys) \<and> (\<forall>y \<in> set ys. y \<in> set xs)" |
|
61166
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
wenzelm
parents:
61115
diff
changeset
|
458 |
proof goal_cases |
60583 | 459 |
{ |
460 |
case 1 |
|
461 |
show ?case (is "?lhs \<longleftrightarrow> ?rhs") |
|
462 |
proof |
|
463 |
show ?rhs if ?lhs |
|
464 |
using that |
|
465 |
by (auto simp add: rhs_def Let_def List.card_set[symmetric] |
|
466 |
card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV |
|
467 |
Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set) |
|
468 |
show ?lhs if ?rhs |
|
469 |
proof - |
|
470 |
have "\<lbrakk> \<forall>y\<in>set xs. y \<notin> set ys; \<forall>x\<in>set ys. x \<notin> set xs \<rbrakk> \<Longrightarrow> set xs \<inter> set ys = {}" by blast |
|
471 |
with that show ?thesis |
|
472 |
by (auto simp add: rhs_def Let_def List.card_set[symmetric] |
|
473 |
card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"] |
|
62390 | 474 |
dest: card_eq_UNIV_imp_eq_UNIV split: if_split_asm) |
60583 | 475 |
qed |
476 |
qed |
|
477 |
} |
|
478 |
moreover |
|
479 |
case 2 |
|
480 |
ultimately show ?case unfolding eq_set_def by blast |
|
481 |
next |
|
482 |
case 3 |
|
483 |
show ?case unfolding eq_set_def List.coset_def by blast |
|
484 |
next |
|
485 |
case 4 |
|
486 |
show ?case unfolding eq_set_def List.coset_def by blast |
|
48051 | 487 |
qed |
488 |
||
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
489 |
end |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
490 |
|
60500 | 491 |
text \<open> |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
492 |
Provide more informative exceptions than Match for non-rewritten cases. |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
493 |
If generated code raises one these exceptions, then a code equation calls |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
494 |
the mentioned operator for an element type that is not an instance of |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
495 |
@{class card_UNIV} and is therefore not implemented via @{term card_UNIV}. |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
496 |
Constrain the element type with sort @{class card_UNIV} to change this. |
60500 | 497 |
\<close> |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
498 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
499 |
lemma card_coset_error [code]: |
53745 | 500 |
"card (List.coset xs) = |
501 |
Code.abort (STR ''card (List.coset _) requires type class instance card_UNIV'') |
|
502 |
(\<lambda>_. card (List.coset xs))" |
|
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
503 |
by(simp) |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
504 |
|
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
505 |
lemma coset_subseteq_set_code [code]: |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
506 |
"List.coset xs \<subseteq> set ys \<longleftrightarrow> |
53745 | 507 |
(if xs = [] \<and> ys = [] then False |
508 |
else Code.abort |
|
509 |
(STR ''subset_eq (List.coset _) (List.set _) requires type class instance card_UNIV'') |
|
510 |
(\<lambda>_. List.coset xs \<subseteq> set ys))" |
|
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
511 |
by simp |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
512 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
513 |
notepad begin \<comment> \<open>test code setup\<close> |
51139
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
514 |
have "List.coset [True] = set [False] \<and> |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
515 |
List.coset [] \<subseteq> List.set [True, False] \<and> |
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents:
51116
diff
changeset
|
516 |
finite (List.coset [True])" |
48062
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
517 |
by eval |
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
518 |
end |
9014e78ccde2
improved code setup for card, finite, subset
Andreas Lochbihler
parents:
48060
diff
changeset
|
519 |
|
48051 | 520 |
end |