src/HOL/Library/Cardinality.thy
author Andreas Lochbihler
Mon, 04 Jun 2012 09:50:57 +0200
changeset 48070 02d64fd40852
parent 48067 9f458b3efb5c
child 48164 e97369f20c30
permissions -rw-r--r--
more sort constraints for FinFun code generation
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
37653
847e95ca9b0a split off Cardinality from Numeral_Type
haftmann
parents: 36350
diff changeset
     5
header {* Cardinality of types *}
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
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44142
diff changeset
     8
imports "~~/src/HOL/Main"
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
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    11
subsection {* Preliminary lemmas *}
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
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    30
lemma finite_range_Some: "finite (range (Some :: 'a \<Rightarrow> 'a option)) = finite (UNIV :: 'a set)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    31
by(auto dest: finite_imageD intro: inj_Some)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    32
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    33
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    34
subsection {* Cardinalities of types *}
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    35
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    36
syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    37
35431
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35362
diff changeset
    38
translations "CARD('t)" => "CONST card (CONST UNIV \<Colon> 't set)"
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    39
42247
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    40
typed_print_translation (advanced) {*
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    41
  let
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    42
    fun card_univ_tr' ctxt _ [Const (@{const_syntax UNIV}, Type (_, [T, _]))] =
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    43
      Syntax.const @{syntax_const "_type_card"} $ Syntax_Phases.term_of_typ ctxt T;
12fe41a92cd5 typed_print_translation: discontinued show_sorts argument;
wenzelm
parents: 42245
diff changeset
    44
  in [(@{const_syntax card}, card_univ_tr')] end
24407
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    45
*}
61b10ffb2549 typed print translation for CARD('a);
huffman
parents: 24406
diff changeset
    46
48058
11a732f7d79f drop redundant sort constraint
Andreas Lochbihler
parents: 48053
diff changeset
    47
lemma card_prod [simp]: "CARD('a \<times> 'b) = CARD('a) * CARD('b)"
26153
b037fd9016fa other UNIV lemmas
haftmann
parents: 25459
diff changeset
    48
  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
    49
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    50
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
    51
unfolding UNIV_Plus_UNIV[symmetric]
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    52
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
    53
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    54
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
    55
by(simp add: card_UNIV_sum)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    56
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    57
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
    58
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    59
  have "(None :: 'a option) \<notin> range Some" by clarsimp
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    60
  thus ?thesis
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    61
    by(simp add: UNIV_option_conv card_eq_0_iff finite_range_Some card_insert_disjoint card_image)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    62
qed
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    63
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    64
lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    65
by(simp add: card_UNIV_option)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    66
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    67
lemma card_UNIV_set: "CARD('a set) = (if CARD('a) = 0 then 0 else 2 ^ CARD('a))"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    68
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
    69
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    70
lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    71
by(simp add: card_UNIV_set)
24332
e3a2b75b1cf9 boolean algebras as locales and numbers as types by Brian Huffman
kleing
parents:
diff changeset
    72
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    73
lemma card_nat [simp]: "CARD(nat) = 0"
44142
8e27e0177518 avoid warnings about duplicate rules
huffman
parents: 42247
diff changeset
    74
  by (simp add: card_eq_0_iff)
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
    75
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    76
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
    77
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    78
  {  assume "0 < CARD('a)" and "0 < CARD('b)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    79
    hence fina: "finite (UNIV :: 'a set)" and finb: "finite (UNIV :: 'b set)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    80
      by(simp_all only: card_ge_0_finite)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    81
    from finite_distinct_list[OF finb] obtain bs 
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    82
      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
    83
    from finite_distinct_list[OF fina] obtain as
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    84
      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
    85
    have cb: "CARD('b) = length bs"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    86
      unfolding bs[symmetric] distinct_card[OF distb] ..
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    87
    have ca: "CARD('a) = length as"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    88
      unfolding as[symmetric] distinct_card[OF dista] ..
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    89
    let ?xs = "map (\<lambda>ys. the o map_of (zip as ys)) (Enum.n_lists (length as) bs)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    90
    have "UNIV = set ?xs"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    91
    proof(rule UNIV_eq_I)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    92
      fix f :: "'a \<Rightarrow> 'b"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    93
      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
    94
        by(auto simp add: map_of_zip_map)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    95
      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
    96
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    97
    moreover have "distinct ?xs" unfolding distinct_map
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    98
    proof(intro conjI distinct_n_lists distb inj_onI)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
    99
      fix xs ys :: "'b list"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   100
      assume xs: "xs \<in> set (Enum.n_lists (length as) bs)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   101
        and ys: "ys \<in> set (Enum.n_lists (length as) bs)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   102
        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
   103
      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
   104
        by(simp_all add: length_n_lists_elem)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   105
      have "map_of (zip as xs) = map_of (zip as ys)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   106
      proof
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   107
        fix x
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   108
        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
   109
          by(simp_all add: map_of_zip_is_Some[symmetric])
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   110
        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
   111
          by(auto dest: fun_cong[where x=x])
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   112
      qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   113
      with dista show "xs = ys" by(simp add: map_of_zip_inject)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   114
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   115
    hence "card (set ?xs) = length ?xs" by(simp only: distinct_card)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   116
    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
   117
    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
   118
  moreover {
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   119
    assume cb: "CARD('b) = 1"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   120
    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
   121
    have eq: "UNIV = {\<lambda>x :: 'a. b ::'b}"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   122
    proof(rule UNIV_eq_I)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   123
      fix x :: "'a \<Rightarrow> 'b"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   124
      { fix y
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   125
        have "x y \<in> UNIV" ..
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   126
        hence "x y = b" unfolding b by simp }
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   127
      thus "x \<in> {\<lambda>x. b}" by(auto)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   128
    qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   129
    have "CARD('a \<Rightarrow> 'b) = 1" unfolding eq by simp }
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   130
  ultimately show ?thesis
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   131
    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
   132
qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   133
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   134
lemma card_nibble: "CARD(nibble) = 16"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   135
unfolding UNIV_nibble by simp
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   136
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   137
lemma card_UNIV_char: "CARD(char) = 256"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   138
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   139
  have "inj (\<lambda>(x, y). Char x y)" by(auto intro: injI)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   140
  thus ?thesis unfolding UNIV_char by(simp add: card_image card_nibble)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   141
qed
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   142
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   143
lemma card_literal: "CARD(String.literal) = 0"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   144
proof -
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   145
  have "inj STR" by(auto intro: injI)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   146
  thus ?thesis by(simp add: type_definition.univ[OF type_definition_literal] card_image infinite_UNIV_listI)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   147
qed
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   148
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   149
subsection {* Classes with at least 1 and 2  *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   150
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   151
text {* Class finite already captures "at least 1" *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   152
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   153
lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)"
29997
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   154
  unfolding neq0_conv [symmetric] by simp
f6756c097c2d number_ring instances for numeral types
huffman
parents: 29629
diff changeset
   155
30001
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   156
lemma one_le_card_finite [simp]: "Suc 0 \<le> CARD('a::finite)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   157
  by (simp add: less_Suc_eq_le [symmetric])
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   158
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   159
text {* Class for cardinality "at least 2" *}
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   160
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   161
class card2 = finite + 
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   162
  assumes two_le_card: "2 \<le> CARD('a)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   163
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   164
lemma one_less_card: "Suc 0 < CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   165
  using two_le_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   166
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   167
lemma one_less_int_card: "1 < int CARD('a::card2)"
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   168
  using one_less_card [where 'a='a] by simp
dd27e16677b2 cleaned up
huffman
parents: 29999
diff changeset
   169
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   170
subsection {* A type class for computing the cardinality of types *}
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   171
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   172
definition is_list_UNIV :: "'a list \<Rightarrow> bool"
48070
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   173
where [code del]: "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
   174
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   175
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
   176
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
   177
   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
   178
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   179
class card_UNIV = 
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   180
  fixes card_UNIV :: "'a itself \<Rightarrow> nat"
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   181
  assumes card_UNIV: "card_UNIV x = CARD('a)"
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   182
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   183
lemma card_UNIV_code [code_unfold]: "CARD('a :: card_UNIV) = card_UNIV TYPE('a)"
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   184
by(simp add: card_UNIV)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   185
48070
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   186
lemma is_list_UNIV_code [code]:
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   187
  "is_list_UNIV (xs :: 'a list) = 
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   188
  (let c = CARD('a :: card_UNIV) in if c = 0 then False else size (remdups xs) = c)"
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   189
by(rule is_list_UNIV_def)
02d64fd40852 more sort constraints for FinFun code generation
Andreas Lochbihler
parents: 48067
diff changeset
   190
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   191
lemma finite_UNIV_conv_card_UNIV [code_unfold]:
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   192
  "finite (UNIV :: 'a :: card_UNIV set) \<longleftrightarrow> card_UNIV TYPE('a) > 0"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   193
by(simp add: card_UNIV card_gt_0_iff)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   194
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   195
subsection {* Instantiations for @{text "card_UNIV"} *}
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   196
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   197
instantiation nat :: card_UNIV begin
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   198
definition "card_UNIV_class.card_UNIV = (\<lambda>a :: nat itself. 0)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   199
instance by intro_classes (simp add: card_UNIV_nat_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   200
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   201
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   202
instantiation int :: card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   203
definition "card_UNIV = (\<lambda>a :: int itself. 0)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   204
instance by intro_classes (simp add: card_UNIV_int_def infinite_UNIV_int)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   205
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   206
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   207
instantiation list :: (type) card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   208
definition "card_UNIV = (\<lambda>a :: 'a list itself. 0)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   209
instance by intro_classes (simp add: card_UNIV_list_def infinite_UNIV_listI)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   210
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   211
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   212
instantiation unit :: card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   213
definition "card_UNIV = (\<lambda>a :: unit itself. 1)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   214
instance by intro_classes (simp add: card_UNIV_unit_def card_UNIV_unit)
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 bool :: card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   218
definition "card_UNIV = (\<lambda>a :: bool itself. 2)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   219
instance by(intro_classes)(simp add: card_UNIV_bool_def card_UNIV_bool)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   220
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   221
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   222
instantiation char :: card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   223
definition "card_UNIV_class.card_UNIV = (\<lambda>a :: char itself. 256)"
48053
9bc78a08ff0a tuned proofs
Andreas Lochbihler
parents: 48052
diff changeset
   224
instance by intro_classes (simp add: card_UNIV_char_def card_UNIV_char)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   225
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   226
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   227
instantiation prod :: (card_UNIV, card_UNIV) card_UNIV begin
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   228
definition "card_UNIV = (\<lambda>a :: ('a \<times> 'b) itself. card_UNIV (TYPE('a)) * card_UNIV (TYPE('b)))"
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   229
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
   230
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   231
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   232
instantiation sum :: (card_UNIV, card_UNIV) card_UNIV begin
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   233
definition "card_UNIV = (\<lambda>a :: ('a + 'b) itself. 
48052
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   234
  let ca = card_UNIV (TYPE('a)); cb = card_UNIV (TYPE('b))
b74766e4c11e tuned instantiations
Andreas Lochbihler
parents: 48051
diff changeset
   235
  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
   236
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
   237
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   238
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   239
instantiation "fun" :: (card_UNIV, card_UNIV) card_UNIV begin
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   240
definition "card_UNIV = (\<lambda>a :: ('a \<Rightarrow> 'b) itself. 
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   241
  let ca = card_UNIV (TYPE('a)); cb = card_UNIV (TYPE('b))
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   242
  in if ca \<noteq> 0 \<and> cb \<noteq> 0 \<or> cb = 1 then cb ^ ca else 0)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   243
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
   244
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   245
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   246
instantiation option :: (card_UNIV) card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   247
definition "card_UNIV = (\<lambda>a :: 'a option itself. 
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   248
  let c = card_UNIV (TYPE('a)) in if c \<noteq> 0 then Suc c else 0)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   249
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
   250
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   251
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   252
instantiation String.literal :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   253
definition "card_UNIV = (\<lambda>a :: String.literal itself. 0)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   254
instance by intro_classes (simp add: card_UNIV_literal_def card_literal)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   255
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   256
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   257
instantiation set :: (card_UNIV) card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   258
definition "card_UNIV = (\<lambda>a :: 'a set itself.
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   259
  let c = card_UNIV (TYPE('a)) in if c = 0 then 0 else 2 ^ c)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   260
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
   261
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   262
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   263
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   264
lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^isub>1]"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   265
by(auto intro: finite_1.exhaust)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   266
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   267
lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^isub>1, finite_2.a\<^isub>2]"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   268
by(auto intro: finite_2.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   269
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   270
lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^isub>1, finite_3.a\<^isub>2, finite_3.a\<^isub>3]"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   271
by(auto intro: finite_3.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   272
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   273
lemma UNIV_finite_4: "UNIV = set [finite_4.a\<^isub>1, finite_4.a\<^isub>2, finite_4.a\<^isub>3, finite_4.a\<^isub>4]"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   274
by(auto intro: finite_4.exhaust)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   275
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   276
lemma UNIV_finite_5:
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   277
  "UNIV = set [finite_5.a\<^isub>1, finite_5.a\<^isub>2, finite_5.a\<^isub>3, finite_5.a\<^isub>4, finite_5.a\<^isub>5]"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   278
by(auto intro: finite_5.exhaust)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   279
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   280
instantiation Enum.finite_1 :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   281
definition "card_UNIV = (\<lambda>a :: Enum.finite_1 itself. 1)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   282
instance by intro_classes (simp add: UNIV_finite_1 card_UNIV_finite_1_def)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   283
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   284
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   285
instantiation Enum.finite_2 :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   286
definition "card_UNIV = (\<lambda>a :: Enum.finite_2 itself. 2)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   287
instance by intro_classes (simp add: UNIV_finite_2 card_UNIV_finite_2_def)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   288
end
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   289
48060
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   290
instantiation Enum.finite_3 :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   291
definition "card_UNIV = (\<lambda>a :: Enum.finite_3 itself. 3)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   292
instance by intro_classes (simp add: UNIV_finite_3 card_UNIV_finite_3_def)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   293
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   294
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   295
instantiation Enum.finite_4 :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   296
definition "card_UNIV = (\<lambda>a :: Enum.finite_4 itself. 4)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   297
instance by intro_classes (simp add: UNIV_finite_4 card_UNIV_finite_4_def)
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   298
end
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   299
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   300
instantiation Enum.finite_5 :: card_UNIV begin
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   301
definition "card_UNIV = (\<lambda>a :: Enum.finite_5 itself. 5)"
1f4d00a7f59f more instantiations for card_UNIV,
Andreas Lochbihler
parents: 48059
diff changeset
   302
instance by intro_classes (simp add: UNIV_finite_5 card_UNIV_finite_5_def)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   303
end
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   304
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   305
subsection {* Code setup for sets *}
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   306
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   307
lemma card_Compl:
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   308
  "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
   309
by (metis Compl_eq_Diff_UNIV card_Diff_subset top_greatest)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   310
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   311
context fixes xs :: "'a :: card_UNIV list"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   312
begin
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   313
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   314
definition card' :: "'a set \<Rightarrow> nat" 
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   315
where [simp, code del, code_abbrev]: "card' = card"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   316
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   317
lemma card'_code [code]:
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   318
  "card' (set xs) = length (remdups xs)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   319
  "card' (List.coset xs) =  card_UNIV TYPE('a) - length (remdups xs)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   320
by(simp_all add: List.card_set card_Compl card_UNIV)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   321
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   322
lemma card'_UNIV [code_unfold]: "card' (UNIV :: 'a :: card_UNIV set) = card_UNIV TYPE('a)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   323
by(simp add: card_UNIV)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   324
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   325
definition finite' :: "'a set \<Rightarrow> bool"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   326
where [simp, code del, code_abbrev]: "finite' = finite"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   327
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   328
lemma finite'_code [code]:
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   329
  "finite' (set xs) \<longleftrightarrow> True"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   330
  "finite' (List.coset xs) \<longleftrightarrow> CARD('a) > 0"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   331
by(simp_all add: card_gt_0_iff)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   332
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   333
definition subset' :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   334
where [simp, code del, code_abbrev]: "subset' = op \<subseteq>"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   335
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   336
lemma subset'_code [code]:
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   337
  "subset' A (List.coset ys) \<longleftrightarrow> (\<forall>y \<in> set ys. y \<notin> A)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   338
  "subset' (set ys) B \<longleftrightarrow> (\<forall>y \<in> set ys. y \<in> B)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   339
  "subset' (List.coset xs) (set ys) \<longleftrightarrow> (let n = CARD('a) in n > 0 \<and> card(set (xs @ ys)) = n)"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   340
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
   341
  (metis finite_compl finite_set rev_finite_subset)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   342
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   343
definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   344
where [simp, code del, code_abbrev]: "eq_set = op ="
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   345
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   346
lemma eq_set_code [code]:
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   347
  fixes ys
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   348
  defines "rhs \<equiv> 
48059
f6ce99d3719b simplify card_UNIV type class,
Andreas Lochbihler
parents: 48058
diff changeset
   349
  let n = CARD('a)
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   350
  in if n = 0 then False else 
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   351
        let xs' = remdups xs; ys' = remdups ys 
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   352
        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')"
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   353
  shows "eq_set (List.coset xs) (set ys) \<longleftrightarrow> rhs" (is ?thesis1)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   354
  and "eq_set (set ys) (List.coset xs) \<longleftrightarrow> rhs" (is ?thesis2)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   355
  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)" (is ?thesis3)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   356
  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)" (is ?thesis4)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   357
proof -
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   358
  show ?thesis1 (is "?lhs \<longleftrightarrow> ?rhs")
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   359
  proof
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   360
    assume ?lhs thus ?rhs
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   361
      by(auto simp add: rhs_def Let_def List.card_set[symmetric] card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   362
  next
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   363
    assume ?rhs
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   364
    moreover 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
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   365
    ultimately show ?lhs
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   366
      by(auto simp add: rhs_def Let_def List.card_set[symmetric] card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"] dest: card_eq_UNIV_imp_eq_UNIV split: split_if_asm)
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   367
  qed
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   368
  thus ?thesis2 unfolding eq_set_def by blast
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   369
  show ?thesis3 ?thesis4 unfolding eq_set_def List.coset_def by blast+
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   370
qed
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   371
48062
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   372
end
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   373
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   374
notepad begin (* test code setup *)
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   375
have "List.coset [True] = set [False] \<and> List.coset [] \<subseteq> List.set [True, False] \<and> finite (List.coset [True])"
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   376
  by eval
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   377
end
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   378
9014e78ccde2 improved code setup for card, finite, subset
Andreas Lochbihler
parents: 48060
diff changeset
   379
hide_const (open) card' finite' subset' eq_set
48051
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   380
53a0df441e20 unify Card_Univ and Cardinality
Andreas Lochbihler
parents: 47221
diff changeset
   381
end