src/HOL/BNF/Basic_BNFs.thy
author traytel
Wed, 13 Nov 2013 10:53:36 +0100
changeset 54421 632be352a5a3
parent 54191 7fba375a7e7d
child 54485 b61b8c9e4cf7
permissions -rw-r--r--
more explicit syntax for defining a bnf
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49507
diff changeset
     1
(*  Title:      HOL/BNF/Basic_BNFs.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Author:     Jasmin Blanchette, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49236
diff changeset
     7
Registration of basic types as bounded natural functors.
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49236
diff changeset
    10
header {* Registration of Basic Types as Bounded Natural Functors *}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
theory Basic_BNFs
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    13
imports BNF_Def
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49310
diff changeset
    16
lemma wpull_id: "wpull UNIV B1 B2 id id id id"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49310
diff changeset
    17
unfolding wpull_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49310
diff changeset
    18
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
lemma ctwo_card_order: "card_order ctwo"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
using Card_order_ctwo by (unfold ctwo_def Field_card_of)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
lemma natLeq_cinfinite: "cinfinite natLeq"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
unfolding cinfinite_def Field_natLeq by (rule nat_infinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    27
lemma wpull_Grp_def: "wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow> Grp B1 f1 OO (Grp B2 f2)\<inverse>\<inverse> \<le> (Grp A p1)\<inverse>\<inverse> OO Grp A p2"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    28
  unfolding wpull_def Grp_def by auto
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    29
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    30
bnf ID: 'a
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    31
  map: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    32
  sets: "\<lambda>x. {x}"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    33
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    34
  rel: "id :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    35
apply (auto simp: Grp_def fun_eq_iff relcompp.simps natLeq_card_order natLeq_cinfinite)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
    37
apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
done
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    40
bnf DEADID: 'a
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    41
  map: "id :: 'a \<Rightarrow> 'a"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    42
  bd: "natLeq +c |UNIV :: 'a set|"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    43
  rel: "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
    44
by (auto simp add: wpull_Grp_def Grp_def
51446
a6ebb12cc003 hide internal constants; tuned proofs
traytel
parents: 49510
diff changeset
    45
  card_order_csum natLeq_card_order card_of_card_order_on
a6ebb12cc003 hide internal constants; tuned proofs
traytel
parents: 49510
diff changeset
    46
  cinfinite_csum natLeq_cinfinite)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    48
definition setl :: "'a + 'b \<Rightarrow> 'a set" where
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    49
"setl x = (case x of Inl z => {z} | _ => {})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    51
definition setr :: "'a + 'b \<Rightarrow> 'b set" where
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    52
"setr x = (case x of Inr z => {z} | _ => {})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    54
lemmas sum_set_defs = setl_def[abs_def] setr_def[abs_def]
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    56
bnf "'a + 'b"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    57
  map: sum_map
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    58
  sets: setl setr
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    59
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    60
  wits: Inl Inr
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
    61
  rel: sum_rel
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
proof -
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
  show "sum_map id id = id" by (rule sum_map.id)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
  fix f1 f2 g1 g2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    66
  show "sum_map (g1 o f1) (g2 o f2) = sum_map g1 g2 o sum_map f1 f2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
    by (rule sum_map.comp[symmetric])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    68
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
  fix x f1 f2 g1 g2
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    70
  assume a1: "\<And>z. z \<in> setl x \<Longrightarrow> f1 z = g1 z" and
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    71
         a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    72
  thus "sum_map f1 f2 x = sum_map g1 g2 x"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
  proof (cases x)
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    74
    case Inl thus ?thesis using a1 by (clarsimp simp: setl_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    75
  next
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    76
    case Inr thus ?thesis using a2 by (clarsimp simp: setr_def)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    77
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    78
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
  fix f1 f2
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    80
  show "setl o sum_map f1 f2 = image f1 o setl"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    81
    by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    83
  fix f1 f2
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    84
  show "setr o sum_map f1 f2 = image f2 o setr"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    85
    by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
  show "card_order natLeq" by (rule natLeq_card_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    89
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
  fix x
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    92
  show "|setl x| \<le>o natLeq"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
    apply (rule ordLess_imp_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
    apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    95
    by (simp add: setl_def split: sum.split)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    96
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
  fix x
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
    98
  show "|setr x| \<le>o natLeq"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
    apply (rule ordLess_imp_ordLeq)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
    apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
   101
    by (simp add: setr_def split: sum.split)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   102
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
  fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
  assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
  hence
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
    pull1: "\<And>b1 b2. \<lbrakk>b1 \<in> B11; b2 \<in> B21; f11 b1 = f21 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A1. p11 a = b1 \<and> p21 a = b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
    and pull2: "\<And>b1 b2. \<lbrakk>b1 \<in> B12; b2 \<in> B22; f12 b1 = f22 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A2. p12 a = b1 \<and> p22 a = b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
    unfolding wpull_def by blast+
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
   109
  show "wpull {x. setl x \<subseteq> A1 \<and> setr x \<subseteq> A2}
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
   110
  {x. setl x \<subseteq> B11 \<and> setr x \<subseteq> B12} {x. setl x \<subseteq> B21 \<and> setr x \<subseteq> B22}
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
  (sum_map f11 f12) (sum_map f21 f22) (sum_map p11 p12) (sum_map p21 p22)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
    (is "wpull ?in ?in1 ?in2 ?mapf1 ?mapf2 ?mapp1 ?mapp2")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
  proof (unfold wpull_def)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
    { fix B1 B2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
      assume *: "B1 \<in> ?in1" "B2 \<in> ?in2" "?mapf1 B1 = ?mapf2 B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
      have "\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
      proof (cases B1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
        case (Inl b1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   119
        { fix b2 assume "B2 = Inr b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
          with Inl *(3) have False by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
        } then obtain b2 where Inl': "B2 = Inl b2" by (cases B2) (simp, blast)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
        with Inl * have "b1 \<in> B11" "b2 \<in> B21" "f11 b1 = f21 b2"
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49450
diff changeset
   123
        by (simp add: setl_def)+
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
        with pull1 obtain a where "a \<in> A1" "p11 a = b1" "p21 a = b2" by blast+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   125
        with Inl Inl' have "Inl a \<in> ?in" "?mapp1 (Inl a) = B1 \<and> ?mapp2 (Inl a) = B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
        by (simp add: sum_set_defs)+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
        thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
      next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
        case (Inr b1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
        { fix b2 assume "B2 = Inl b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   131
          with Inr *(3) have False by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
        } then obtain b2 where Inr': "B2 = Inr b2" by (cases B2) (simp, blast)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
        with Inr * have "b1 \<in> B12" "b2 \<in> B22" "f12 b1 = f22 b2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
        by (simp add: sum_set_defs)+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   135
        with pull2 obtain a where "a \<in> A2" "p12 a = b1" "p22 a = b2" by blast+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
        with Inr Inr' have "Inr a \<in> ?in" "?mapp1 (Inr a) = B1 \<and> ?mapp2 (Inr a) = B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
        by (simp add: sum_set_defs)+
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
        thus ?thesis by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   139
      qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
    }
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
    thus "\<forall>B1 B2. B1 \<in> ?in1 \<and> B2 \<in> ?in2 \<and> ?mapf1 B1 = ?mapf2 B2 \<longrightarrow>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   142
      (\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2)" by fastforce
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
  qed
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   144
next
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   145
  fix R S
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   146
  show "sum_rel R S =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   147
        (Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (sum_map fst fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   148
        Grp {x. setl x \<subseteq> Collect (split R) \<and> setr x \<subseteq> Collect (split S)} (sum_map snd snd)"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   149
  unfolding setl_def setr_def sum_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   150
  by (fastforce split: sum.splits)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   151
qed (auto simp: sum_set_defs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   152
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
definition fsts :: "'a \<times> 'b \<Rightarrow> 'a set" where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   154
"fsts x = {fst x}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   156
definition snds :: "'a \<times> 'b \<Rightarrow> 'b set" where
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
"snds x = {snd x}"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def]
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   161
bnf "'a \<times> 'b"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   162
  map: map_pair
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   163
  sets: fsts snds
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   164
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   165
  rel: prod_rel
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
proof (unfold prod_set_defs)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
  show "map_pair id id = id" by (rule map_pair.id)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
  fix f1 f2 g1 g2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  show "map_pair (g1 o f1) (g2 o f2) = map_pair g1 g2 o map_pair f1 f2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
    by (rule map_pair.comp[symmetric])
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
  fix x f1 f2 g1 g2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
  assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
  thus "map_pair f1 f2 x = map_pair g1 g2 x" by (cases x) simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
  fix f1 f2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
  show "(\<lambda>x. {fst x}) o map_pair f1 f2 = image f1 o (\<lambda>x. {fst x})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   179
    by (rule ext, unfold o_apply) simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   180
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   181
  fix f1 f2
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   182
  show "(\<lambda>x. {snd x}) o map_pair f1 f2 = image f2 o (\<lambda>x. {snd x})"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   183
    by (rule ext, unfold o_apply) simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   184
next
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   185
  show "card_order natLeq" by (rule natLeq_card_order)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
next
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   187
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
  fix x
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   190
  show "|{fst x}| \<le>o natLeq"
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   191
    by (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq finite.emptyI finite_insert)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   192
next
52635
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   193
  fix x
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   194
  show "|{snd x}| \<le>o natLeq"
4f84b730c489 got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
traytel
parents: 52545
diff changeset
   195
    by (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq finite.emptyI finite_insert)
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   196
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   197
  fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   198
  assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   199
  thus "wpull {x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   200
    {x. {fst x} \<subseteq> B11 \<and> {snd x} \<subseteq> B12} {x. {fst x} \<subseteq> B21 \<and> {snd x} \<subseteq> B22}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   201
   (map_pair f11 f12) (map_pair f21 f22) (map_pair p11 p12) (map_pair p21 p22)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   202
    unfolding wpull_def by simp fast
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   203
next
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   204
  fix R S
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   205
  show "prod_rel R S =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   206
        (Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair fst fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   207
        Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair snd snd)"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   208
  unfolding prod_set_defs prod_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   209
  by auto
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 53026
diff changeset
   210
qed
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   211
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   212
(* Categorical version of pullback: *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   213
lemma wpull_cat:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   214
assumes p: "wpull A B1 B2 f1 f2 p1 p2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   215
and c: "f1 o q1 = f2 o q2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   216
and r: "range q1 \<subseteq> B1" "range q2 \<subseteq> B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   217
obtains h where "range h \<subseteq> A \<and> q1 = p1 o h \<and> q2 = p2 o h"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   218
proof-
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   219
  have *: "\<forall>d. \<exists>a \<in> A. p1 a = q1 d & p2 a = q2 d"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   220
  proof safe
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   221
    fix d
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   222
    have "f1 (q1 d) = f2 (q2 d)" using c unfolding comp_def[abs_def] by (rule fun_cong)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   223
    moreover
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   224
    have "q1 d : B1" "q2 d : B2" using r unfolding image_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   225
    ultimately show "\<exists>a \<in> A. p1 a = q1 d \<and> p2 a = q2 d"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   226
      using p unfolding wpull_def by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   228
  then obtain h where "!! d. h d \<in> A & p1 (h d) = q1 d & p2 (h d) = q2 d" by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   229
  thus ?thesis using that by fastforce
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   230
qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   231
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   232
lemma card_of_bounded_range:
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   233
  "|{f :: 'd \<Rightarrow> 'a. range f \<subseteq> B}| \<le>o |Func (UNIV :: 'd set) B|" (is "|?LHS| \<le>o |?RHS|")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   234
proof -
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 51893
diff changeset
   235
  let ?f = "\<lambda>f. %x. if f x \<in> B then f x else undefined"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   236
  have "inj_on ?f ?LHS" unfolding inj_on_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   237
  proof (unfold fun_eq_iff, safe)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   238
    fix g :: "'d \<Rightarrow> 'a" and f :: "'d \<Rightarrow> 'a" and x
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   239
    assume "range f \<subseteq> B" "range g \<subseteq> B" and eq: "\<forall>x. ?f f x = ?f g x"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   240
    hence "f x \<in> B" "g x \<in> B" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   241
    with eq have "Some (f x) = Some (g x)" by metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   242
    thus "f x = g x" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   243
  qed
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   244
  moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Func_def by fastforce
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   245
  ultimately show ?thesis using card_of_ordLeq by fast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   246
qed
54191
7fba375a7e7d removed junk
traytel
parents: 54189
diff changeset
   247
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   248
bnf "'a \<Rightarrow> 'b"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   249
  map: "op \<circ>"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   250
  sets: range
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   251
  bd: "natLeq +c |UNIV :: 'a set|"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54191
diff changeset
   252
  rel: "fun_rel op ="
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   253
proof
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   254
  fix f show "id \<circ> f = id f" by simp
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   255
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   256
  fix f g show "op \<circ> (g \<circ> f) = op \<circ> g \<circ> op \<circ> f"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   257
  unfolding comp_def[abs_def] ..
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   258
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   259
  fix x f g
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   260
  assume "\<And>z. z \<in> range x \<Longrightarrow> f z = g z"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   261
  thus "f \<circ> x = g \<circ> x" by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   262
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   263
  fix f show "range \<circ> op \<circ> f = op ` f \<circ> range"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   264
  unfolding image_def comp_def[abs_def] by auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   265
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   266
  show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)")
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   267
  apply (rule card_order_csum)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   268
  apply (rule natLeq_card_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   269
  by (rule card_of_card_order_on)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   270
(*  *)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   271
  show "cinfinite (natLeq +c ?U)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   272
    apply (rule cinfinite_csum)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   273
    apply (rule disjI1)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   274
    by (rule natLeq_cinfinite)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   275
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   276
  fix f :: "'d => 'a"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   277
  have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   278
  also have "?U \<le>o natLeq +c ?U"  by (rule ordLeq_csum2) (rule card_of_Card_order)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   279
  finally show "|range f| \<le>o natLeq +c ?U" .
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   280
next
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   281
  fix A B1 B2 f1 f2 p1 p2 assume p: "wpull A B1 B2 f1 f2 p1 p2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   282
  show "wpull {h. range h \<subseteq> A} {g1. range g1 \<subseteq> B1} {g2. range g2 \<subseteq> B2}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   283
    (op \<circ> f1) (op \<circ> f2) (op \<circ> p1) (op \<circ> p2)"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   284
  unfolding wpull_def
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   285
  proof safe
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   286
    fix g1 g2 assume r: "range g1 \<subseteq> B1" "range g2 \<subseteq> B2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   287
    and c: "f1 \<circ> g1 = f2 \<circ> g2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   288
    show "\<exists>h \<in> {h. range h \<subseteq> A}. p1 \<circ> h = g1 \<and> p2 \<circ> h = g2"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   289
    using wpull_cat[OF p c r] by simp metis
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   290
  qed
49453
ff0e540d8758 add rel as first-class citizen of BNF
blanchet
parents: 49451
diff changeset
   291
next
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49455
diff changeset
   292
  fix R
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   293
  show "fun_rel op = R =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   294
        (Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   295
         Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   296
  unfolding fun_rel_def Grp_def fun_eq_iff relcompp.simps conversep.simps  subset_iff image_iff
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   297
  by auto (force, metis pair_collapse)
54189
c0186a0d8cb3 define a trivial nonemptiness witness if none is provided
traytel
parents: 53026
diff changeset
   298
qed
54191
7fba375a7e7d removed junk
traytel
parents: 54189
diff changeset
   299
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   300
end