src/HOL/Library/Cardinality.thy
author nipkow
Tue, 15 Jan 2019 21:31:20 +0100
changeset 69663 41ff40bf1530
parent 69593 3dda49e08b9d
child 71174 7fac205dd737
permissions -rw-r--r--
moved and renamed class
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37653
847e95ca9b0a split off Cardinality from Numeral_Type
haftmann
parents: 36350
diff changeset
     1
(*  Title:      HOL/Library/Cardinality.thy
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
     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
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     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
847e95ca9b0a split off Cardinality from Numeral_Type
haftmann
parents: 36350
diff changeset
     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
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    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
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    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
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60679
diff changeset
    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
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    37
print_translation \<open>
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    38
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
    39
    fun card_univ_tr' ctxt [Const (\<^const_syntax>\<open>UNIV\<close>, Type (_, [T]))] =
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
    40
      Syntax.const \<^syntax_const>\<open>_type_card\<close> $ Syntax_Phases.term_of_typ ctxt T
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
    41
  in [(\<^const_syntax>\<open>card\<close>, card_univ_tr')] end
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    42
\<close>
24407
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    43
48058
11a732f7d79f drop redundant sort constraint
Andreas Lochbihler
parents: 48053
diff changeset
    44
lemma card_prod [simp]: "CARD('a \<times> 'b) = CARD('a) * CARD('b)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    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
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    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)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    48
unfolding UNIV_Plus_UNIV[symmetric]
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    49
by(auto simp add: card_eq_0_iff card_Plus simp del: UNIV_Plus_UNIV)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    50
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    51
lemma card_sum [simp]: "CARD('a + 'b) = CARD('a::finite) + CARD('b::finite)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    52
by(simp add: card_UNIV_sum)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    53
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    54
lemma card_UNIV_option: "CARD('a option) = (if CARD('a) = 0 then 0 else CARD('a) + 1)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    55
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    56
  have "(None :: 'a option) \<notin> range Some" by clarsimp
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    57
  thus ?thesis
53191
14ab2f821e1d tuned proofs -- fewer warnings;
wenzelm
parents: 53015
diff changeset
    58
    by (simp add: UNIV_option_conv card_eq_0_iff finite_range_Some card_image)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    59
qed
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    60
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    61
lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    62
by(simp add: card_UNIV_option)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    63
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    64
lemma card_UNIV_set: "CARD('a set) = (if CARD('a) = 0 then 0 else 2 ^ CARD('a))"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 68028
diff changeset
    65
by(simp add: card_eq_0_iff card_Pow flip: Pow_UNIV)
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    66
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    67
lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    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
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    70
lemma card_nat [simp]: "CARD(nat) = 0"
44142
8e27e0177518 avoid warnings about duplicate rules
huffman
parents: 42247
diff changeset
    71
  by (simp add: card_eq_0_iff)
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    72
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    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)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    74
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    75
  {  assume "0 < CARD('a)" and "0 < CARD('b)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    76
    hence fina: "finite (UNIV :: 'a set)" and finb: "finite (UNIV :: 'b set)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    77
      by(simp_all only: card_ge_0_finite)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    78
    from finite_distinct_list[OF finb] obtain bs 
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    79
      where bs: "set bs = (UNIV :: 'b set)" and distb: "distinct bs" by blast
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    80
    from finite_distinct_list[OF fina] obtain as
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    81
      where as: "set as = (UNIV :: 'a set)" and dista: "distinct as" by blast
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    82
    have cb: "CARD('b) = length bs"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    83
      unfolding bs[symmetric] distinct_card[OF distb] ..
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    84
    have ca: "CARD('a) = length as"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    85
      unfolding as[symmetric] distinct_card[OF dista] ..
67091
1393c2340eec more symbols;
wenzelm
parents: 62597
diff changeset
    86
    let ?xs = "map (\<lambda>ys. the \<circ> map_of (zip as ys)) (List.n_lists (length as) bs)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    87
    have "UNIV = set ?xs"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    88
    proof(rule UNIV_eq_I)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    89
      fix f :: "'a \<Rightarrow> 'b"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    90
      from as have "f = the \<circ> map_of (zip as (map f as))"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    91
        by(auto simp add: map_of_zip_map)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    92
      thus "f \<in> set ?xs" using bs by(auto simp add: set_n_lists)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    93
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    94
    moreover have "distinct ?xs" unfolding distinct_map
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    95
    proof(intro conjI distinct_n_lists distb inj_onI)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    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
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    99
        and eq: "the \<circ> map_of (zip as xs) = the \<circ> map_of (zip as ys)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   100
      from xs ys have [simp]: "length xs = length as" "length ys = length as"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   101
        by(simp_all add: length_n_lists_elem)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   102
      have "map_of (zip as xs) = map_of (zip as ys)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   103
      proof
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   104
        fix x
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   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"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   106
          by(simp_all add: map_of_zip_is_Some[symmetric])
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   107
        with eq show "map_of (zip as xs) x = map_of (zip as ys) x"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   108
          by(auto dest: fun_cong[where x=x])
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   109
      qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   110
      with dista show "xs = ys" by(simp add: map_of_zip_inject)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   111
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   112
    hence "card (set ?xs) = length ?xs" by(simp only: distinct_card)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   113
    moreover have "length ?xs = length bs ^ length as" by(simp add: length_n_lists)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   114
    ultimately have "CARD('a \<Rightarrow> 'b) = CARD('b) ^ CARD('a)" using cb ca by simp }
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   115
  moreover {
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   116
    assume cb: "CARD('b) = 1"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   117
    then obtain b where b: "UNIV = {b :: 'b}" by(auto simp add: card_Suc_eq)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   118
    have eq: "UNIV = {\<lambda>x :: 'a. b ::'b}"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   119
    proof(rule UNIV_eq_I)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   120
      fix x :: "'a \<Rightarrow> 'b"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   121
      { fix y
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   122
        have "x y \<in> UNIV" ..
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   123
        hence "x y = b" unfolding b by simp }
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   124
      thus "x \<in> {\<lambda>x. b}" by(auto)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   125
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   126
    have "CARD('a \<Rightarrow> 'b) = 1" unfolding eq by simp }
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   127
  ultimately show ?thesis
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   128
    by(auto simp del: One_nat_def)(auto simp add: card_eq_0_iff dest: finite_fun_UNIVD2 finite_fun_UNIVD1)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   129
qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   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
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   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
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   145
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   146
subsection \<open>Classes with at least 1 and 2\<close>
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   147
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   148
text \<open>Class finite already captures "at least 1"\<close>
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   149
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   150
lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   151
  unfolding neq0_conv [symmetric] by simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   152
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   153
lemma one_le_card_finite [simp]: "Suc 0 \<le> CARD('a::finite)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   154
  by (simp add: less_Suc_eq_le [symmetric])
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   155
69663
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   156
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   157
class CARD_1 =
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   158
  assumes CARD_1: "CARD ('a) = 1"
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   159
begin
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   160
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   161
subclass finite
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   162
proof
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   163
  from CARD_1 show "finite (UNIV :: 'a set)"
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   164
    by (auto intro!: card_ge_0_finite)
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   165
qed
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   166
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   167
end
41ff40bf1530 moved and renamed class
nipkow
parents: 69593
diff changeset
   168
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   169
text \<open>Class for cardinality "at least 2"\<close>
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   170
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   171
class card2 = finite + 
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   172
  assumes two_le_card: "2 \<le> CARD('a)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   173
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   174
lemma one_less_card: "Suc 0 < CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   175
  using two_le_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   176
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   177
lemma one_less_int_card: "1 < int CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   178
  using one_less_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   179
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   180
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   181
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
   182
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   183
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
   184
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   185
class finite_UNIV = 
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   186
  fixes finite_UNIV :: "('a, bool) phantom"
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   187
  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
   188
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   189
lemma finite_UNIV_code [code_unfold]:
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   190
  "finite (UNIV :: 'a :: finite_UNIV set)
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   191
  \<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
   192
by(simp add: finite_UNIV)
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   193
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   194
subsection \<open>A type class for computing the cardinality of types\<close>
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   195
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   196
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
   197
where "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)"
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   198
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   199
lemma is_list_UNIV_iff: "is_list_UNIV xs \<longleftrightarrow> set xs = UNIV"
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   200
by(auto simp add: is_list_UNIV_def Let_def card_eq_0_iff List.card_set[symmetric] 
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   201
   dest: subst[where P="finite", OF _ finite_set] card_eq_UNIV_imp_eq_UNIV)
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   202
48164
e97369f20c30 change card_UNIV from itself to phantom type to avoid unnecessary closures in generated code
Andreas Lochbihler
parents: 48070
diff changeset
   203
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
   204
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   205
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
   206
  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
   207
  assumes card_UNIV: "card_UNIV = Phantom('a) CARD('a)"
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   208
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61166
diff changeset
   209
subsection \<open>Instantiations for \<open>card_UNIV\<close>\<close>
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   210
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   211
instantiation nat :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   212
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
   213
definition "card_UNIV = Phantom(nat) 0"
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   214
instance by intro_classes (simp_all add: finite_UNIV_nat_def card_UNIV_nat_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   215
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   216
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   217
instantiation int :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   218
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
   219
definition "card_UNIV = Phantom(int) 0"
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   220
instance by intro_classes (simp_all add: card_UNIV_int_def finite_UNIV_int_def infinite_UNIV_int)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   221
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   222
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51139
diff changeset
   223
instantiation natural :: card_UNIV begin
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51139
diff changeset
   224
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
   225
definition "card_UNIV = Phantom(natural) 0"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   226
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   227
  by standard
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   228
    (auto simp add: finite_UNIV_natural_def card_UNIV_natural_def card_eq_0_iff
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   229
      type_definition.univ [OF type_definition_natural] natural_eq_iff
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   230
      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
   231
end
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51139
diff changeset
   232
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51139
diff changeset
   233
instantiation integer :: card_UNIV begin
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51139
diff changeset
   234
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
   235
definition "card_UNIV = Phantom(integer) 0"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   236
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   237
  by standard
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   238
    (auto simp add: finite_UNIV_integer_def card_UNIV_integer_def card_eq_0_iff
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   239
      type_definition.univ [OF type_definition_integer] infinite_UNIV_int
ade12ef2773c tuned proofs;
wenzelm
parents: 60583
diff changeset
   240
      dest!: finite_imageD intro: inj_onI)
48165
d07a0b9601aa instantiate card_UNIV with nibble and code_numeral
Andreas Lochbihler
parents: 48164
diff changeset
   241
end
d07a0b9601aa instantiate card_UNIV with nibble and code_numeral
Andreas Lochbihler
parents: 48164
diff changeset
   242
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   243
instantiation list :: (type) card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   244
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
   245
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
   246
instance by intro_classes (simp_all add: card_UNIV_list_def finite_UNIV_list_def infinite_UNIV_listI)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   247
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   248
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   249
instantiation unit :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   250
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
   251
definition "card_UNIV = Phantom(unit) 1"
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   252
instance by intro_classes (simp_all add: card_UNIV_unit_def finite_UNIV_unit_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   253
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   254
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   255
instantiation bool :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   256
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
   257
definition "card_UNIV = Phantom(bool) 2"
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   258
instance by(intro_classes)(simp_all add: card_UNIV_bool_def finite_UNIV_bool_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   259
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   260
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   261
instantiation char :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   262
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
   263
definition "card_UNIV = Phantom(char) 256"
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   264
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
   265
end
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   266
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   267
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
   268
definition "finite_UNIV = Phantom('a \<times> 'b) 
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   269
  (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
   270
instance by intro_classes (simp add: finite_UNIV_prod_def finite_UNIV finite_prod)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   271
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   272
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   273
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
   274
definition "card_UNIV = Phantom('a \<times> 'b) 
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   275
  (of_phantom (card_UNIV :: 'a card_UNIV) * of_phantom (card_UNIV :: 'b card_UNIV))"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   276
instance by intro_classes (simp add: card_UNIV_prod_def card_UNIV)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   277
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   278
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   279
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
   280
definition "finite_UNIV = Phantom('a + 'b)
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   281
  (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
   282
instance
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 68028
diff changeset
   283
  by intro_classes (simp add: finite_UNIV_sum_def finite_UNIV)
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   284
end
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   285
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   286
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
   287
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
   288
  (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
   289
       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
   290
   in if ca \<noteq> 0 \<and> cb \<noteq> 0 then ca + cb else 0)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   291
instance by intro_classes (auto simp add: card_UNIV_sum_def card_UNIV card_UNIV_sum)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   292
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   293
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   294
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
   295
definition "finite_UNIV = Phantom('a \<Rightarrow> 'b)
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   296
  (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
   297
   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
   298
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   299
  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
   300
end
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   301
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   302
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
   303
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
   304
  (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
   305
       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
   306
   in if ca \<noteq> 0 \<and> cb \<noteq> 0 \<or> cb = 1 then cb ^ ca else 0)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   307
instance by intro_classes (simp add: card_UNIV_fun_def card_UNIV Let_def card_fun)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   308
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   309
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   310
instantiation option :: (finite_UNIV) finite_UNIV begin
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   311
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
   312
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
   313
end
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   314
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   315
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
   316
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
   317
  (let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c \<noteq> 0 then Suc c else 0)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   318
instance by intro_classes (simp add: card_UNIV_option_def card_UNIV card_UNIV_option)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   319
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   320
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   321
instantiation String.literal :: card_UNIV begin
48176
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   322
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
   323
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
   324
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   325
  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
   326
end
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   327
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   328
instantiation set :: (finite_UNIV) finite_UNIV begin
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   329
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
   330
instance by intro_classes (simp add: finite_UNIV_set_def finite_UNIV Finite_Set.finite_set)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   331
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   332
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   333
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
   334
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
   335
  (let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c = 0 then 0 else 2 ^ c)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   336
instance by intro_classes (simp add: card_UNIV_set_def card_UNIV_set card_UNIV)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   337
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   338
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52147
diff changeset
   339
lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^sub>1]"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   340
by(auto intro: finite_1.exhaust)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   341
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52147
diff changeset
   342
lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^sub>1, finite_2.a\<^sub>2]"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   343
by(auto intro: finite_2.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   344
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52147
diff changeset
   345
lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^sub>1, finite_3.a\<^sub>2, finite_3.a\<^sub>3]"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   346
by(auto intro: finite_3.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   347
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52147
diff changeset
   348
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
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   349
by(auto intro: finite_4.exhaust)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   350
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   351
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
   352
  "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
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   353
by(auto intro: finite_5.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   354
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   355
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
   356
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
   357
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
   358
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   359
  by intro_classes (simp_all add: UNIV_finite_1 card_UNIV_finite_1_def finite_UNIV_finite_1_def)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   360
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   361
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   362
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
   363
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
   364
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
   365
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   366
  by intro_classes (simp_all add: UNIV_finite_2 card_UNIV_finite_2_def finite_UNIV_finite_2_def)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   367
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   368
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   369
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
   370
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
   371
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
   372
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   373
  by intro_classes (simp_all add: UNIV_finite_3 card_UNIV_finite_3_def finite_UNIV_finite_3_def)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   374
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   375
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   376
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
   377
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
   378
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
   379
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   380
  by intro_classes (simp_all add: UNIV_finite_4 card_UNIV_finite_4_def finite_UNIV_finite_4_def)
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   381
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   382
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   383
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
   384
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
   385
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
   386
instance
3d9c1ddb9f47 new type class for computing finiteness of types with instantiations
Andreas Lochbihler
parents: 48165
diff changeset
   387
  by intro_classes (simp_all add: UNIV_finite_5 card_UNIV_finite_5_def finite_UNIV_finite_5_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   388
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   389
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   390
subsection \<open>Code setup for sets\<close>
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   391
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   392
text \<open>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   393
  Implement \<^term>\<open>CARD('a)\<close> via \<^term>\<open>card_UNIV\<close> and provide
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   394
  implementations for \<^term>\<open>finite\<close>, \<^term>\<open>card\<close>, \<^term>\<open>(\<subseteq>)\<close>, 
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   395
  and \<^term>\<open>(=)\<close>if the calling context already provides \<^class>\<open>finite_UNIV\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   396
  and \<^class>\<open>card_UNIV\<close> instances. If we implemented the latter
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   397
  always via \<^term>\<open>card_UNIV\<close>, we would require instances of essentially all 
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
   398
  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
   399
  possibly slow dictionary constructions.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   400
\<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
   401
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   402
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   403
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   404
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   405
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
   406
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
   407
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
   408
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
   409
  "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
   410
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
   411
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
   412
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
   413
  "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
   414
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
   415
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   416
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
   417
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   418
lemma card_Compl:
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   419
  "finite A \<Longrightarrow> card (- A) = card (UNIV :: 'a set) - card (A :: 'a set)"
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   420
by (metis Compl_eq_Diff_UNIV card_Diff_subset top_greatest)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   421
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
   422
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
   423
begin
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   424
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   425
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
   426
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
   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 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
   429
  "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
   430
  "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
   431
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
   432
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
   433
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
   434
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
   435
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
   436
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
   437
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   438
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
   439
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
   440
 
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
   441
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
   442
  "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
   443
  "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
   444
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
   445
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
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   447
qualified definition subset' :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   448
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
   449
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
   450
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
   451
  "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
   452
  "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
   453
  "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
   454
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
   455
  (metis finite_compl finite_set rev_finite_subset)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   456
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 61076
diff changeset
   457
qualified definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67091
diff changeset
   458
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
   459
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
   460
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
   461
  fixes ys
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   462
  defines "rhs \<equiv> 
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   463
  let n = CARD('a)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   464
  in if n = 0 then False else 
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   465
        let xs' = remdups xs; ys' = remdups ys 
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   466
        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
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   467
  shows "eq_set (List.coset xs) (set ys) \<longleftrightarrow> rhs"
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   468
  and "eq_set (set ys) (List.coset xs) \<longleftrightarrow> rhs"
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   469
  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)"
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   470
  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
   471
proof goal_cases
60583
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   472
  {
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   473
    case 1
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   474
    show ?case (is "?lhs \<longleftrightarrow> ?rhs")
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   475
    proof
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   476
      show ?rhs if ?lhs
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   477
        using that
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   478
        by (auto simp add: rhs_def Let_def List.card_set[symmetric]
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   479
          card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   480
          Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set)
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   481
      show ?lhs if ?rhs
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   482
      proof - 
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   483
        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
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   484
        with that show ?thesis
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   485
          by (auto simp add: rhs_def Let_def List.card_set[symmetric]
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   486
            card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"]
62390
842917225d56 more canonical names
nipkow
parents: 61585
diff changeset
   487
            dest: card_eq_UNIV_imp_eq_UNIV split: if_split_asm)
60583
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   488
      qed
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   489
    qed
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   490
  }
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   491
  moreover
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   492
  case 2
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   493
  ultimately show ?case unfolding eq_set_def by blast
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   494
next
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   495
  case 3
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   496
  show ?case unfolding eq_set_def List.coset_def by blast
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   497
next
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   498
  case 4
a645a0e6d790 tuned proofs;
wenzelm
parents: 60500
diff changeset
   499
  show ?case unfolding eq_set_def List.coset_def by blast
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   500
qed
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   501
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
   502
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
   503
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   504
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
   505
  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
   506
  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
   507
  the mentioned operator for an element type that is not an instance of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   508
  \<^class>\<open>card_UNIV\<close> and is therefore not implemented via \<^term>\<open>card_UNIV\<close>.
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68406
diff changeset
   509
  Constrain the element type with sort \<^class>\<open>card_UNIV\<close> to change this.
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   510
\<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
   511
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
lemma card_coset_error [code]:
53745
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   513
  "card (List.coset xs) = 
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   514
   Code.abort (STR ''card (List.coset _) requires type class instance card_UNIV'')
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   515
     (\<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
   516
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
   517
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
   518
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
   519
  "List.coset xs \<subseteq> set ys \<longleftrightarrow> 
53745
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   520
  (if xs = [] \<and> ys = [] then False 
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   521
   else Code.abort
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   522
     (STR ''subset_eq (List.coset _) (List.set _) requires type class instance card_UNIV'')
788730ab7da4 prefer Code.abort over code_abort
Andreas Lochbihler
parents: 53191
diff changeset
   523
     (\<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
   524
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
   525
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67399
diff changeset
   526
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
   527
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
   528
      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
   529
      finite (List.coset [True])"
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   530
  by eval
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   531
end
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   532
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   533
end