src/HOL/BNF/More_BNFs.thy
author traytel
Sun, 07 Jul 2013 10:24:00 +0200
changeset 52545 d2ad6eae514f
parent 52544 0c4b140cff00
child 52635 4f84b730c489
permissions -rw-r--r--
Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
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/More_BNFs.thy
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     4
    Author:     Andreas Lochbihler, Karlsruhe Institute of Technology
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     5
    Author:     Jasmin Blanchette, TU Muenchen
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     6
    Copyright   2012
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     7
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     8
Registration of various types as bounded natural functors.
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
     9
*)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    10
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    11
header {* Registration of Various Types as Bounded Natural Functors *}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    12
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    13
theory More_BNFs
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    14
imports
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    15
  BNF_LFP
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    16
  BNF_GFP
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    17
  "~~/src/HOL/Quotient_Examples/Lift_FSet"
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    18
  "~~/src/HOL/Library/Multiset"
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50027
diff changeset
    19
  Countable_Type
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    20
begin
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    21
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    22
lemma option_rec_conv_option_case: "option_rec = option_case"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    23
by (simp add: fun_eq_iff split: option.split)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    24
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51782
diff changeset
    25
bnf Option.map [Option.set] "\<lambda>_::'a option. natLeq" ["None"] option_rel
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    26
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    27
  show "Option.map id = id" by (simp add: fun_eq_iff Option.map_def split: option.split)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    28
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    29
  fix f g
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    30
  show "Option.map (g \<circ> f) = Option.map g \<circ> Option.map f"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    31
    by (auto simp add: fun_eq_iff Option.map_def split: option.split)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    32
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    33
  fix f g x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    34
  assume "\<And>z. z \<in> Option.set x \<Longrightarrow> f z = g z"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    35
  thus "Option.map f x = Option.map g x"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    36
    by (simp cong: Option.map_cong)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    37
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    38
  fix f
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    39
  show "Option.set \<circ> Option.map f = op ` f \<circ> Option.set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    40
    by fastforce
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    41
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    42
  show "card_order natLeq" by (rule natLeq_card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    43
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    44
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    45
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    46
  fix x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    47
  show "|Option.set x| \<le>o natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    48
    by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    49
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    50
  fix A
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    51
  have unfold: "{x. Option.set x \<subseteq> A} = Some ` A \<union> {None}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    52
    by (auto simp add: option_rec_conv_option_case Option.set_def split: option.split_asm)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    53
  show "|{x. Option.set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    54
    apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    55
    apply (rule card_of_ordIso_subst[OF unfold])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    56
    apply (rule ordLeq_transitive)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    57
    apply (rule Un_csum)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    58
    apply (rule ordLeq_transitive)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    59
    apply (rule csum_mono)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    60
    apply (rule card_of_image)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    61
    apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    62
    apply (rule single_cone)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    63
    apply (rule cone_ordLeq_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    64
    apply (rule ordLeq_cexp1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    65
    apply (simp_all add: natLeq_cinfinite natLeq_Card_order cinfinite_not_czero Card_order_csum)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    66
    done
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    67
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    68
  fix A B1 B2 f1 f2 p1 p2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    69
  assume wpull: "wpull A B1 B2 f1 f2 p1 p2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    70
  show "wpull {x. Option.set x \<subseteq> A} {x. Option.set x \<subseteq> B1} {x. Option.set x \<subseteq> B2}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    71
    (Option.map f1) (Option.map f2) (Option.map p1) (Option.map p2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    72
    (is "wpull ?A ?B1 ?B2 ?f1 ?f2 ?p1 ?p2")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    73
    unfolding wpull_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    74
  proof (intro strip, elim conjE)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    75
    fix b1 b2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    76
    assume "b1 \<in> ?B1" "b2 \<in> ?B2" "?f1 b1 = ?f2 b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    77
    thus "\<exists>a \<in> ?A. ?p1 a = b1 \<and> ?p2 a = b2" using wpull
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    78
      unfolding wpull_def by (cases b2) (auto 4 5)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    79
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    80
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    81
  fix z
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    82
  assume "z \<in> Option.set None"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    83
  thus False by simp
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
    84
next
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
    85
  fix R
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    86
  show "option_rel R =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    87
        (Grp {x. Option.set x \<subseteq> Collect (split R)} (Option.map fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    88
         Grp {x. Option.set x \<subseteq> Collect (split R)} (Option.map snd)"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    89
  unfolding option_rel_unfold Grp_def relcompp.simps conversep.simps fun_eq_iff prod.cases
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
    90
  by (auto simp: trans[OF eq_commute option_map_is_None] trans[OF eq_commute option_map_eq_Some]
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    91
           split: option.splits)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    92
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    93
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    94
lemma card_of_list_in:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    95
  "|{xs. set xs \<subseteq> A}| \<le>o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \<le>o |?RHS|")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    96
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    97
  let ?f = "%xs. %i. if i < length xs \<and> set xs \<subseteq> A then Some (nth xs i) else None"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    98
  have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    99
  proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   100
    fix xs :: "'a list" and ys :: "'a list"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   101
    assume su: "set xs \<subseteq> A" "set ys \<subseteq> A" and eq: "\<forall>i. ?f xs i = ?f ys i"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   102
    hence *: "length xs = length ys"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   103
    by (metis linorder_cases option.simps(2) order_less_irrefl)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   104
    thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   105
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   106
  moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Pfunc_def by fastforce
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   107
  ultimately show ?thesis using card_of_ordLeq by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   108
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   109
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   110
lemma list_in_empty: "A = {} \<Longrightarrow> {x. set x \<subseteq> A} = {[]}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   111
by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   112
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   113
lemma card_of_Func: "|Func A B| =o |B| ^c |A|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   114
unfolding cexp_def Field_card_of by (rule card_of_refl)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   115
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   116
lemma not_emp_czero_notIn_ordIso_Card_order:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   117
"A \<noteq> {} \<Longrightarrow> ( |A|, czero) \<notin> ordIso \<and> Card_order |A|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   118
  apply (rule conjI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   119
  apply (metis Field_card_of czeroE)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   120
  by (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   121
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   122
lemma list_in_bd: "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   123
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   124
  fix A :: "'a set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   125
  show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   126
  proof (cases "A = {}")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   127
    case False thus ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   128
      apply -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   129
      apply (rule ordLeq_transitive)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   130
      apply (rule card_of_list_in)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   131
      apply (rule ordLeq_transitive)
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
   132
      apply (erule card_of_Pfunc_Pow_Func_option)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   133
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   134
      apply (rule Times_cprod)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   135
      apply (rule cprod_cinfinite_bound)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   136
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   137
      apply (rule Pow_cexp_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   138
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   139
      apply (rule cexp_cong2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   140
      apply (rule card_of_nat)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   141
      apply (rule Card_order_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   142
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   143
      apply (rule cexp_mono1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   144
      apply (rule ordLeq_csum2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   145
      apply (rule Card_order_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   146
      apply (rule natLeq_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   147
      apply (rule ordIso_ordLeq_trans)
52545
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
   148
      apply (rule card_of_Func_option_Func)
d2ad6eae514f Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents: 52544
diff changeset
   149
      apply (rule ordIso_ordLeq_trans)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   150
      apply (rule card_of_Func)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   151
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   152
      apply (rule cexp_cong2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   153
      apply (rule card_of_nat)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   154
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   155
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   156
      apply (rule cexp_mono1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   157
      apply (rule ordLeq_csum1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   158
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   159
      apply (rule natLeq_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   160
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   161
      apply (rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   162
      apply (rule Cinfinite_cexp)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   163
      apply (rule ordLeq_csum2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   164
      apply (rule Card_order_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   165
      apply (rule conjI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   166
      apply (rule natLeq_cinfinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   167
      by (rule natLeq_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   168
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   169
    case True thus ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   170
      apply -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   171
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   172
      apply (rule card_of_ordIso_subst)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   173
      apply (erule list_in_empty)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   174
      apply (rule ordIso_ordLeq_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   175
      apply (rule single_cone)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   176
      apply (rule cone_ordLeq_cexp)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   177
      apply (rule ordLeq_transitive)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   178
      apply (rule cone_ordLeq_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   179
      apply (rule ordLeq_csum2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   180
      by (rule Card_order_ctwo)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   181
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   182
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   183
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   184
lemma wpull_map:
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   185
  assumes "wpull A B1 B2 f1 f2 p1 p2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   186
  shows "wpull {x. set x \<subseteq> A} {x. set x \<subseteq> B1} {x. set x \<subseteq> B2} (map f1) (map f2) (map p1) (map p2)"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   187
    (is "wpull ?A ?B1 ?B2 _ _ _ _")
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   188
proof (unfold wpull_def)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   189
  { fix as bs assume *: "as \<in> ?B1" "bs \<in> ?B2" "map f1 as = map f2 bs"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   190
    hence "length as = length bs" by (metis length_map)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   191
    hence "\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs" using *
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   192
    proof (induct as bs rule: list_induct2)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   193
      case (Cons a as b bs)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   194
      hence "a \<in> B1" "b \<in> B2" "f1 a = f2 b" by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   195
      with assms obtain z where "z \<in> A" "p1 z = a" "p2 z = b" unfolding wpull_def by blast
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   196
      moreover
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   197
      from Cons obtain zs where "zs \<in> ?A" "map p1 zs = as" "map p2 zs = bs" by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   198
      ultimately have "z # zs \<in> ?A" "map p1 (z # zs) = a # as \<and> map p2 (z # zs) = b # bs" by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   199
      thus ?case by (rule_tac x = "z # zs" in bexI)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   200
    qed simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   201
  }
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   202
  thus "\<forall>as bs. as \<in> ?B1 \<and> bs \<in> ?B2 \<and> map f1 as = map f2 bs \<longrightarrow>
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   203
    (\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs)" by blast
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   204
qed
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   205
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51782
diff changeset
   206
bnf map [set] "\<lambda>_::'a list. natLeq" ["[]"]
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   207
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   208
  show "map id = id" by (rule List.map.id)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   209
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   210
  fix f g
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   211
  show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   212
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   213
  fix x f g
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   214
  assume "\<And>z. z \<in> set x \<Longrightarrow> f z = g z"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   215
  thus "map f x = map g x" by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   216
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   217
  fix f
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   218
  show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   219
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   220
  show "card_order natLeq" by (rule natLeq_card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   221
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   222
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   223
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   224
  fix x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   225
  show "|set x| \<le>o natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   226
    apply (rule ordLess_imp_ordLeq)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   227
    apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   228
    unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   229
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   230
  fix A :: "'a set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   231
  show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   232
qed (simp add: wpull_map)+
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   233
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   234
(* Finite sets *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   235
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   236
definition fset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" where
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   237
"fset_rel R a b \<longleftrightarrow>
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   238
 (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and>
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   239
 (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   240
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   241
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   242
lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   243
  by (rule f_the_inv_into_f[unfolded inj_on_def])
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   244
    (transfer, simp,
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   245
     transfer, metis Collect_finite_eq_lists lists_UNIV mem_Collect_eq)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   246
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   247
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   248
lemma fset_rel_aux:
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   249
"(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow>
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   250
 ((Grp {a. fset a \<subseteq> {(a, b). R a b}} (fmap fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   251
  Grp {a. fset a \<subseteq> {(a, b). R a b}} (fmap snd)) a b" (is "?L = ?R")
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   252
proof
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   253
  assume ?L
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   254
  def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'")
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   255
  have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   256
  hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   257
  show ?R unfolding Grp_def relcompp.simps conversep.simps
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   258
  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   259
    from * show "a = fmap fst R'" using conjunct1[OF `?L`]
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   260
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   261
    from * show "b = fmap snd R'" using conjunct2[OF `?L`]
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   262
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   263
  qed (auto simp add: *)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   264
next
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   265
  assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   266
  apply (simp add: subset_eq Ball_def)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   267
  apply (rule conjI)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   268
  apply (transfer, clarsimp, metis snd_conv)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   269
  by (transfer, clarsimp, metis fst_conv)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   270
qed
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   271
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   272
lemma abs_fset_rep_fset[simp]: "abs_fset (rep_fset x) = x"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   273
  by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format])
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   274
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   275
lemma wpull_image:
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   276
  assumes "wpull A B1 B2 f1 f2 p1 p2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   277
  shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   278
unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   279
  fix Y1 Y2 assume Y1: "Y1 \<subseteq> B1" and Y2: "Y2 \<subseteq> B2" and EQ: "f1 ` Y1 = f2 ` Y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   280
  def X \<equiv> "{a \<in> A. p1 a \<in> Y1 \<and> p2 a \<in> Y2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   281
  show "\<exists>X\<subseteq>A. p1 ` X = Y1 \<and> p2 ` X = Y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   282
  proof (rule exI[of _ X], intro conjI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   283
    show "p1 ` X = Y1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   284
    proof
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   285
      show "Y1 \<subseteq> p1 ` X"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   286
      proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   287
        fix y1 assume y1: "y1 \<in> Y1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   288
        then obtain y2 where y2: "y2 \<in> Y2" and eq: "f1 y1 = f2 y2" using EQ by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   289
        then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   290
        using assms y1 Y1 Y2 unfolding wpull_def by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   291
        thus "y1 \<in> p1 ` X" unfolding X_def using y1 y2 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   292
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   293
    qed(unfold X_def, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   294
    show "p2 ` X = Y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   295
    proof
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   296
      show "Y2 \<subseteq> p2 ` X"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   297
      proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   298
        fix y2 assume y2: "y2 \<in> Y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   299
        then obtain y1 where y1: "y1 \<in> Y1" and eq: "f1 y1 = f2 y2" using EQ by force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   300
        then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   301
        using assms y2 Y1 Y2 unfolding wpull_def by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   302
        thus "y2 \<in> p2 ` X" unfolding X_def using y1 y2 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   303
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   304
    qed(unfold X_def, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   305
  qed(unfold X_def, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   306
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   307
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   308
lemma wpull_fmap:
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   309
  assumes "wpull A B1 B2 f1 f2 p1 p2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   310
  shows "wpull {x. fset x \<subseteq> A} {x. fset x \<subseteq> B1} {x. fset x \<subseteq> B2}
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   311
              (fmap f1) (fmap f2) (fmap p1) (fmap p2)"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   312
unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   313
  fix y1 y2
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   314
  assume Y1: "fset y1 \<subseteq> B1" and Y2: "fset y2 \<subseteq> B2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   315
  assume "fmap f1 y1 = fmap f2 y2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   316
  hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" by transfer simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   317
  with Y1 Y2 obtain X where X: "X \<subseteq> A" and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   318
    using wpull_image[OF assms] unfolding wpull_def Pow_def
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   319
    by (auto elim!: allE[of _ "fset y1"] allE[of _ "fset y2"])
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   320
  have "\<forall> y1' \<in> fset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   321
  then obtain q1 where q1: "\<forall> y1' \<in> fset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   322
  have "\<forall> y2' \<in> fset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   323
  then obtain q2 where q2: "\<forall> y2' \<in> fset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   324
  def X' \<equiv> "q1 ` (fset y1) \<union> q2 ` (fset y2)"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   325
  have X': "X' \<subseteq> A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   326
  using X Y1 Y2 q1 q2 unfolding X'_def by auto
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   327
  have fX': "finite X'" unfolding X'_def by transfer simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   328
  then obtain x where X'eq: "X' = fset x" by transfer (metis finite_list)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   329
  show "\<exists>x. fset x \<subseteq> A \<and> fmap p1 x = y1 \<and> fmap p2 x = y2"
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   330
     using X' Y1 Y2 by (auto simp: X'eq intro!: exI[of _ "x"]) (transfer, simp)+
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   331
qed
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   332
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51782
diff changeset
   333
bnf fmap [fset] "\<lambda>_::'a fset. natLeq" ["{||}"] fset_rel
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   334
apply -
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   335
          apply transfer' apply simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   336
         apply transfer' apply simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   337
        apply transfer apply force
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   338
       apply transfer apply force
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   339
      apply (rule natLeq_card_order)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   340
     apply (rule natLeq_cinfinite)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   341
    apply transfer apply (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq finite_set)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   342
   apply (rule ordLeq_transitive[OF surj_imp_ordLeq[of _ abs_fset] list_in_bd])
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   343
   apply (auto simp: fset_def intro!: image_eqI[of _ abs_fset]) []
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   344
  apply (erule wpull_fmap)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   345
 apply (simp add: Grp_def relcompp.simps conversep.simps fun_eq_iff fset_rel_def fset_rel_aux) 
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   346
apply transfer apply simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   347
done
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   348
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
   349
lemmas [simp] = fset.map_comp' fset.map_id' fset.set_map'
51371
197ad6b8f763 some simp rules for fset
traytel
parents: 50144
diff changeset
   350
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
   351
lemma fset_rel_fset: "set_rel \<chi> (fset A1) (fset A2) = fset_rel \<chi> A1 A2"
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   352
  unfolding fset_rel_def set_rel_def by auto 
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
   353
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   354
(* Countable sets *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   355
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   356
lemma card_of_countable_sets_range:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   357
fixes A :: "'a set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   358
shows "|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |{f::nat \<Rightarrow> 'a. range f \<subseteq> A}|"
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50027
diff changeset
   359
apply(rule card_of_ordLeqI[of from_nat_into]) using inj_on_from_nat_into
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   360
unfolding inj_on_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   361
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   362
lemma card_of_countable_sets_Func:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   363
"|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |A| ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   364
using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   365
unfolding cexp_def Field_natLeq Field_card_of
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   366
by (rule ordLeq_ordIso_trans)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   367
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   368
lemma ordLeq_countable_subsets:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   369
"|A| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   370
apply (rule card_of_ordLeqI[of "\<lambda> a. {a}"]) unfolding inj_on_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   371
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   372
lemma finite_countable_subset:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   373
"finite {X. X \<subseteq> A \<and> countable X} \<longleftrightarrow> finite A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   374
apply default
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   375
 apply (erule contrapos_pp)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   376
 apply (rule card_of_ordLeq_infinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   377
 apply (rule ordLeq_countable_subsets)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   378
 apply assumption
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   379
apply (rule finite_Collect_conjI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   380
apply (rule disjI1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   381
by (erule finite_Collect_subsets)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   382
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   383
lemma card_of_countable_sets:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   384
"|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   385
(is "|?L| \<le>o _")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   386
proof(cases "finite A")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   387
  let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   388
  case True hence "finite ?L" by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   389
  moreover have "infinite ?R"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   390
  apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   391
  ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   392
  apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   393
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   394
  case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   395
  hence "|{X. X \<subseteq> A \<and> countable X}| =o |{X. X \<subseteq> A \<and> countable X} - {{}}|"
52544
0c4b140cff00 tuned spelling
traytel
parents: 51893
diff changeset
   396
  by (intro card_of_infinite_diff_finite finite.emptyI finite.insertI ordIso_symmetric)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   397
     (unfold finite_countable_subset)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   398
  also have "|{X. X \<subseteq> A \<and> countable X} - {{}}| \<le>o |A| ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   399
  using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   400
  also have "|A| ^c natLeq \<le>o ( |A| +c ctwo) ^c natLeq"
51782
84e7225f5ab6 removed unnecessary assumptions in some theorems about cardinal exponentiation
traytel
parents: 51766
diff changeset
   401
  apply(rule cexp_mono1)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   402
    apply(rule ordLeq_csum1, rule card_of_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   403
    by (rule natLeq_Card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   404
  finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   405
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   406
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   407
lemma rcset_to_rcset: "countable A \<Longrightarrow> rcset (the_inv rcset A) = A"
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   408
apply (rule f_the_inv_into_f)
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   409
unfolding inj_on_def rcset_inj using rcset_surj by auto
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   410
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   411
lemma Collect_Int_Times:
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   412
"{(x, y). R x y} \<inter> A \<times> B = {(x, y). R x y \<and> x \<in> A \<and> y \<in> B}"
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   413
by auto
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   414
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
   415
lemma rcset_map': "rcset (cIm f x) = f ` rcset x"
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   416
unfolding cIm_def[abs_def] by simp
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   417
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
   418
definition cset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a cset \<Rightarrow> 'b cset \<Rightarrow> bool" where
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
   419
"cset_rel R a b \<longleftrightarrow>
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   420
 (\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and>
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   421
 (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t)"
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   422
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
   423
lemma cset_rel_aux:
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   424
"(\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and> (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t) \<longleftrightarrow>
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   425
 ((Grp {x. rcset x \<subseteq> {(a, b). R a b}} (cIm fst))\<inverse>\<inverse> OO
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   426
          Grp {x. rcset x \<subseteq> {(a, b). R a b}} (cIm snd)) a b" (is "?L = ?R")
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   427
proof
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   428
  assume ?L
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   429
  def R' \<equiv> "the_inv rcset (Collect (split R) \<inter> (rcset a \<times> rcset b))"
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   430
  (is "the_inv rcset ?L'")
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   431
  have "countable ?L'" by auto
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   432
  hence *: "rcset R' = ?L'" unfolding R'_def using fset_to_fset by (intro rcset_to_rcset)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   433
  show ?R unfolding Grp_def relcompp.simps conversep.simps
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   434
  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   435
    have "rcset a = fst ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?A")
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   436
    using conjunct1[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   437
    hence "a = acset ?A" by (metis acset_rcset)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   438
    thus "a = cIm fst R'" unfolding cIm_def * by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   439
    have "rcset b = snd ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?B")
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   440
    using conjunct2[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   441
    hence "b = acset ?B" by (metis acset_rcset)
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   442
    thus "b = cIm snd R'" unfolding cIm_def * by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   443
  qed (auto simp add: *)
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   444
next
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   445
  assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   446
  apply (simp add: subset_eq Ball_def)
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   447
  apply (rule conjI)
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
   448
  apply (clarsimp, metis (lifting, no_types) rcset_map' image_iff surjective_pairing)
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   449
  apply (clarsimp)
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
   450
  by (metis Domain.intros Range.simps rcset_map' fst_eq_Domain snd_eq_Range)
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   451
qed
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   452
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51782
diff changeset
   453
bnf cIm [rcset] "\<lambda>_::'a cset. natLeq" ["cEmp"] cset_rel
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   454
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   455
  show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   456
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   457
  fix f g show "cIm (g \<circ> f) = cIm g \<circ> cIm f"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   458
  unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   459
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   460
  fix C f g assume eq: "\<And>a. a \<in> rcset C \<Longrightarrow> f a = g a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   461
  thus "cIm f C = cIm g C"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   462
  unfolding cIm_def[abs_def] unfolding image_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   463
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   464
  fix f show "rcset \<circ> cIm f = op ` f \<circ> rcset" unfolding cIm_def[abs_def] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   465
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   466
  show "card_order natLeq" by (rule natLeq_card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   467
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   468
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   469
next
50144
885deccc264e renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
hoelzl
parents: 50027
diff changeset
   470
  fix C show "|rcset C| \<le>o natLeq" using rcset unfolding countable_card_le_natLeq .
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   471
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   472
  fix A :: "'a set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   473
  have "|{Z. rcset Z \<subseteq> A}| \<le>o |acset ` {X. X \<subseteq> A \<and> countable X}|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   474
  apply(rule card_of_mono1) unfolding Pow_def image_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   475
  proof (rule Collect_mono, clarsimp)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   476
    fix x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   477
    assume "rcset x \<subseteq> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   478
    hence "rcset x \<subseteq> A \<and> countable (rcset x) \<and> x = acset (rcset x)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   479
    using acset_rcset[of x] rcset[of x] by force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   480
    thus "\<exists>y \<subseteq> A. countable y \<and> x = acset y" by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   481
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   482
  also have "|acset ` {X. X \<subseteq> A \<and> countable X}| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   483
  using card_of_image .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   484
  also have "|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   485
  using card_of_countable_sets .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   486
  finally show "|{Z. rcset Z \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   487
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   488
  fix A B1 B2 f1 f2 p1 p2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   489
  assume wp: "wpull A B1 B2 f1 f2 p1 p2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   490
  show "wpull {x. rcset x \<subseteq> A} {x. rcset x \<subseteq> B1} {x. rcset x \<subseteq> B2}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   491
              (cIm f1) (cIm f2) (cIm p1) (cIm p2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   492
  unfolding wpull_def proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   493
    fix y1 y2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   494
    assume Y1: "rcset y1 \<subseteq> B1" and Y2: "rcset y2 \<subseteq> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   495
    assume "cIm f1 y1 = cIm f2 y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   496
    hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   497
    unfolding cIm_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   498
    with Y1 Y2 obtain X where X: "X \<subseteq> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   499
    and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   500
    using wpull_image[OF wp] unfolding wpull_def Pow_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   501
    unfolding Bex_def mem_Collect_eq apply -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   502
    apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   503
    have "\<forall> y1' \<in> rcset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   504
    then obtain q1 where q1: "\<forall> y1' \<in> rcset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   505
    have "\<forall> y2' \<in> rcset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   506
    then obtain q2 where q2: "\<forall> y2' \<in> rcset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   507
    def X' \<equiv> "q1 ` (rcset y1) \<union> q2 ` (rcset y2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   508
    have X': "X' \<subseteq> A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   509
    using X Y1 Y2 q1 q2 unfolding X'_def by fast+
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   510
    have fX': "countable X'" unfolding X'_def by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   511
    then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   512
    show "\<exists>x\<in>{x. rcset x \<subseteq> A}. cIm p1 x = y1 \<and> cIm p2 x = y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   513
    apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   514
  qed
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   515
next
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   516
  fix R
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   517
  show "cset_rel R =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   518
        (Grp {x. rcset x \<subseteq> Collect (split R)} (cIm fst))\<inverse>\<inverse> OO Grp {x. rcset x \<subseteq> Collect (split R)} (cIm snd)"
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   519
  unfolding cset_rel_def[abs_def] cset_rel_aux by simp
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   520
qed (unfold cEmp_def, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   521
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   522
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   523
(* Multisets *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   524
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   525
lemma setsum_gt_0_iff:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   526
fixes f :: "'a \<Rightarrow> nat" assumes "finite A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   527
shows "setsum f A > 0 \<longleftrightarrow> (\<exists> a \<in> A. f a > 0)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   528
(is "?L \<longleftrightarrow> ?R")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   529
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   530
  have "?L \<longleftrightarrow> \<not> setsum f A = 0" by fast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   531
  also have "... \<longleftrightarrow> (\<exists> a \<in> A. f a \<noteq> 0)" using assms by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   532
  also have "... \<longleftrightarrow> ?R" by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   533
  finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   534
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   535
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   536
(*   *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   537
definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'b \<Rightarrow> nat" where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   538
"mmap h f b = setsum f {a. h a = b \<and> f a > 0}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   539
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   540
lemma mmap_id: "mmap id = id"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   541
proof (rule ext)+
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   542
  fix f a show "mmap id f a = id f a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   543
  proof(cases "f a = 0")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   544
    case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   545
    hence 1: "{aa. aa = a \<and> 0 < f aa} = {a}" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   546
    show ?thesis by (simp add: mmap_def id_apply 1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   547
  qed(unfold mmap_def, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   548
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   549
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   550
lemma inj_on_setsum_inv:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   551
assumes f: "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   552
and 1: "(0::nat) < setsum f {a. h a = b' \<and> 0 < f a}" (is "0 < setsum f ?A'")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   553
and 2: "{a. h a = b \<and> 0 < f a} = {a. h a = b' \<and> 0 < f a}" (is "?A = ?A'")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   554
shows "b = b'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   555
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   556
  have "finite ?A'" using f unfolding multiset_def by auto
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49878
diff changeset
   557
  hence "?A' \<noteq> {}" using 1 by (auto simp add: setsum_gt_0_iff)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   558
  thus ?thesis using 2 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   559
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   560
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   561
lemma mmap_comp:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   562
fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   563
assumes f: "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   564
shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   565
unfolding mmap_def[abs_def] comp_def proof(rule ext)+
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   566
  fix c :: 'c
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   567
  let ?A = "{a. h2 (h1 a) = c \<and> 0 < f a}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   568
  let ?As = "\<lambda> b. {a. h1 a = b \<and> 0 < f a}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   569
  let ?B = "{b. h2 b = c \<and> 0 < setsum f (?As b)}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   570
  have 0: "{?As b | b.  b \<in> ?B} = ?As ` ?B" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   571
  have "\<And> b. finite (?As b)" using f unfolding multiset_def by simp
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49878
diff changeset
   572
  hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   573
  hence A: "?A = \<Union> {?As b | b.  b \<in> ?B}" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   574
  have "setsum f ?A = setsum (setsum f) {?As b | b.  b \<in> ?B}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   575
  unfolding A apply(rule setsum_Union_disjoint)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   576
  using f unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   577
  also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 ..
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   578
  also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   579
  unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   580
  also have "... = setsum (\<lambda> b. setsum f (?As b)) ?B" unfolding comp_def ..
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   581
  finally show "setsum f ?A = setsum (\<lambda> b. setsum f (?As b)) ?B" .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   582
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   583
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   584
lemma mmap_comp1:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   585
fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   586
assumes "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   587
shows "mmap (\<lambda> a. h2 (h1 a)) f = mmap h2 (mmap h1 f)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   588
using mmap_comp[OF assms] unfolding comp_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   589
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   590
lemma mmap:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   591
assumes "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   592
shows "mmap h f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   593
using assms unfolding mmap_def[abs_def] multiset_def proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   594
  assume fin: "finite {a. 0 < f a}"  (is "finite ?A")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   595
  show "finite {b. 0 < setsum f {a. h a = b \<and> 0 < f a}}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   596
  (is "finite {b. 0 < setsum f (?As b)}")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   597
  proof- let ?B = "{b. 0 < setsum f (?As b)}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   598
    have "\<And> b. finite (?As b)" using assms unfolding multiset_def by simp
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49878
diff changeset
   599
    hence B: "?B = {b. ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   600
    hence "?B \<subseteq> h ` ?A" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   601
    thus ?thesis using finite_surj[OF fin] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   602
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   603
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   604
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   605
lemma mmap_cong:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   606
assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   607
shows "mmap f (count M) = mmap g (count M)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   608
using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   609
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   610
abbreviation supp where "supp f \<equiv> {a. f a > 0}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   611
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   612
lemma mmap_image_comp:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   613
assumes f: "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   614
shows "(supp o mmap h) f = (image h o supp) f"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   615
unfolding mmap_def[abs_def] comp_def proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   616
  have "\<And> b. finite {a. h a = b \<and> 0 < f a}" (is "\<And> b. finite (?As b)")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   617
  using f unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   618
  thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}"
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49878
diff changeset
   619
  by (auto simp add:  setsum_gt_0_iff)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   620
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   621
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   622
lemma mmap_image:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   623
assumes f: "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   624
shows "supp (mmap h f) = h ` (supp f)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   625
using mmap_image_comp[OF assms] unfolding comp_def .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   626
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   627
lemma set_of_Abs_multiset:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   628
assumes f: "f \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   629
shows "set_of (Abs_multiset f) = supp f"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   630
using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   631
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   632
lemma supp_count:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   633
"supp (count M) = set_of M"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   634
using assms unfolding set_of_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   635
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   636
lemma multiset_of_surj:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   637
"multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   638
proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   639
  fix M assume M: "set_of M \<subseteq> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   640
  obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   641
  hence "set as \<subseteq> A" using M by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   642
  thus "M \<in> multiset_of ` {as. set as \<subseteq> A}" using eq by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   643
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   644
  show "\<And>x xa xb. \<lbrakk>set xa \<subseteq> A; xb \<in> set_of (multiset_of xa)\<rbrakk> \<Longrightarrow> xb \<in> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   645
  by (erule set_mp) (unfold set_of_multiset_of)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   646
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   647
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   648
lemma card_of_set_of:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   649
"|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   650
apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   651
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   652
lemma nat_sum_induct:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   653
assumes "\<And>n1 n2. (\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> phi m1 m2) \<Longrightarrow> phi n1 n2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   654
shows "phi (n1::nat) (n2::nat)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   655
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   656
  let ?chi = "\<lambda> n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   657
  have "?chi (n1,n2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   658
  apply(induct rule: measure_induct[of "\<lambda> n1n2. fst n1n2 + snd n1n2" ?chi])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   659
  using assms by (metis fstI sndI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   660
  thus ?thesis by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   661
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   662
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   663
lemma matrix_count:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   664
fixes ct1 ct2 :: "nat \<Rightarrow> nat"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   665
assumes "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   666
shows
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   667
"\<exists> ct. (\<forall> i1 \<le> n1. setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ct1 i1) \<and>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   668
       (\<forall> i2 \<le> n2. setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ct2 i2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   669
(is "?phi ct1 ct2 n1 n2")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   670
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   671
  have "\<forall> ct1 ct2 :: nat \<Rightarrow> nat.
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   672
        setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   673
  proof(induct rule: nat_sum_induct[of
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   674
"\<lambda> n1 n2. \<forall> ct1 ct2 :: nat \<Rightarrow> nat.
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   675
     setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"],
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   676
      clarify)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   677
  fix n1 n2 :: nat and ct1 ct2 :: "nat \<Rightarrow> nat"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   678
  assume IH: "\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   679
                \<forall> dt1 dt2 :: nat \<Rightarrow> nat.
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   680
                setsum dt1 {..<Suc m1} = setsum dt2 {..<Suc m2} \<longrightarrow> ?phi dt1 dt2 m1 m2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   681
  and ss: "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   682
  show "?phi ct1 ct2 n1 n2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   683
  proof(cases n1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   684
    case 0 note n1 = 0
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   685
    show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   686
    proof(cases n2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   687
      case 0 note n2 = 0
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   688
      let ?ct = "\<lambda> i1 i2. ct2 0"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   689
      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   690
    next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   691
      case (Suc m2) note n2 = Suc
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   692
      let ?ct = "\<lambda> i1 i2. ct2 i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   693
      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   694
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   695
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   696
    case (Suc m1) note n1 = Suc
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   697
    show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   698
    proof(cases n2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   699
      case 0 note n2 = 0
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   700
      let ?ct = "\<lambda> i1 i2. ct1 i1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   701
      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   702
    next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   703
      case (Suc m2) note n2 = Suc
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   704
      show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   705
      proof(cases "ct1 n1 \<le> ct2 n2")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   706
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   707
        def dt2 \<equiv> "\<lambda> i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   708
        have "setsum ct1 {..<Suc m1} = setsum dt2 {..<Suc n2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   709
        unfolding dt2_def using ss n1 True by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   710
        hence "?phi ct1 dt2 m1 n2" using IH[of m1 n2] n1 by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   711
        then obtain dt where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   712
        1: "\<And> i1. i1 \<le> m1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc n2} = ct1 i1" and
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   713
        2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc m1} = dt2 i2" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   714
        let ?ct = "\<lambda> i1 i2. if i1 = n1 then (if i2 = n2 then ct1 n1 else 0)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   715
                                       else dt i1 i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   716
        show ?thesis apply(rule exI[of _ ?ct])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   717
        using n1 n2 1 2 True unfolding dt2_def by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   718
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   719
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   720
        hence False: "ct2 n2 < ct1 n1" by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   721
        def dt1 \<equiv> "\<lambda> i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   722
        have "setsum dt1 {..<Suc n1} = setsum ct2 {..<Suc m2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   723
        unfolding dt1_def using ss n2 False by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   724
        hence "?phi dt1 ct2 n1 m2" using IH[of n1 m2] n2 by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   725
        then obtain dt where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   726
        1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc m2} = dt1 i1" and
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   727
        2: "\<And> i2. i2 \<le> m2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc n1} = ct2 i2" by force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   728
        let ?ct = "\<lambda> i1 i2. if i2 = n2 then (if i1 = n1 then ct2 n2 else 0)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   729
                                       else dt i1 i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   730
        show ?thesis apply(rule exI[of _ ?ct])
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   731
        using n1 n2 1 2 False unfolding dt1_def by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   732
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   733
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   734
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   735
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   736
  thus ?thesis using assms by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   737
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   738
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   739
definition
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   740
"inj2 u B1 B2 \<equiv>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   741
 \<forall> b1 b1' b2 b2'. {b1,b1'} \<subseteq> B1 \<and> {b2,b2'} \<subseteq> B2 \<and> u b1 b2 = u b1' b2'
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   742
                  \<longrightarrow> b1 = b1' \<and> b2 = b2'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   743
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   744
lemma matrix_setsum_finite:
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   745
assumes B1: "B1 \<noteq> {}" "finite B1" and B2: "B2 \<noteq> {}" "finite B2" and u: "inj2 u B1 B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   746
and ss: "setsum N1 B1 = setsum N2 B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   747
shows "\<exists> M :: 'a \<Rightarrow> nat.
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   748
            (\<forall> b1 \<in> B1. setsum (\<lambda> b2. M (u b1 b2)) B2 = N1 b1) \<and>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   749
            (\<forall> b2 \<in> B2. setsum (\<lambda> b1. M (u b1 b2)) B1 = N2 b2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   750
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   751
  obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   752
  then obtain e1 where e1: "bij_betw e1 {..<Suc n1} B1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   753
  using ex_bij_betw_finite_nat[OF B1(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   754
  hence e1_inj: "inj_on e1 {..<Suc n1}" and e1_surj: "e1 ` {..<Suc n1} = B1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   755
  unfolding bij_betw_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   756
  def f1 \<equiv> "inv_into {..<Suc n1} e1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   757
  have f1: "bij_betw f1 B1 {..<Suc n1}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   758
  and f1e1[simp]: "\<And> i1. i1 < Suc n1 \<Longrightarrow> f1 (e1 i1) = i1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   759
  and e1f1[simp]: "\<And> b1. b1 \<in> B1 \<Longrightarrow> e1 (f1 b1) = b1" unfolding f1_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   760
  apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   761
  by (metis e1_surj f_inv_into_f)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   762
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   763
  obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   764
  then obtain e2 where e2: "bij_betw e2 {..<Suc n2} B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   765
  using ex_bij_betw_finite_nat[OF B2(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   766
  hence e2_inj: "inj_on e2 {..<Suc n2}" and e2_surj: "e2 ` {..<Suc n2} = B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   767
  unfolding bij_betw_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   768
  def f2 \<equiv> "inv_into {..<Suc n2} e2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   769
  have f2: "bij_betw f2 B2 {..<Suc n2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   770
  and f2e2[simp]: "\<And> i2. i2 < Suc n2 \<Longrightarrow> f2 (e2 i2) = i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   771
  and e2f2[simp]: "\<And> b2. b2 \<in> B2 \<Longrightarrow> e2 (f2 b2) = b2" unfolding f2_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   772
  apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   773
  by (metis e2_surj f_inv_into_f)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   774
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   775
  let ?ct1 = "N1 o e1"  let ?ct2 = "N2 o e2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   776
  have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   777
  unfolding setsum_reindex[OF e1_inj, symmetric] setsum_reindex[OF e2_inj, symmetric]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   778
  e1_surj e2_surj using ss .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   779
  obtain ct where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   780
  ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   781
  ct2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ?ct2 i2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   782
  using matrix_count[OF ss] by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   783
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   784
  def A \<equiv> "{u b1 b2 | b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   785
  have "\<forall> a \<in> A. \<exists> b1b2 \<in> B1 <*> B2. u (fst b1b2) (snd b1b2) = a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   786
  unfolding A_def Ball_def mem_Collect_eq by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   787
  then obtain h1h2 where h12:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   788
  "\<And>a. a \<in> A \<Longrightarrow> u (fst (h1h2 a)) (snd (h1h2 a)) = a \<and> h1h2 a \<in> B1 <*> B2" by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   789
  def h1 \<equiv> "fst o h1h2"  def h2 \<equiv> "snd o h1h2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   790
  have h12[simp]: "\<And>a. a \<in> A \<Longrightarrow> u (h1 a) (h2 a) = a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   791
                  "\<And> a. a \<in> A \<Longrightarrow> h1 a \<in> B1"  "\<And> a. a \<in> A \<Longrightarrow> h2 a \<in> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   792
  using h12 unfolding h1_def h2_def by force+
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   793
  {fix b1 b2 assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   794
   hence inA: "u b1 b2 \<in> A" unfolding A_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   795
   hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   796
   moreover have "h1 (u b1 b2) \<in> B1" "h2 (u b1 b2) \<in> B2" using inA by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   797
   ultimately have "h1 (u b1 b2) = b1 \<and> h2 (u b1 b2) = b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   798
   using u b1 b2 unfolding inj2_def by fastforce
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   799
  }
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   800
  hence h1[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h1 (u b1 b2) = b1" and
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   801
        h2[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h2 (u b1 b2) = b2" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   802
  def M \<equiv> "\<lambda> a. ct (f1 (h1 a)) (f2 (h2 a))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   803
  show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   804
  apply(rule exI[of _ M]) proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   805
    fix b1 assume b1: "b1 \<in> B1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   806
    hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   807
    by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   808
    have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   809
    unfolding e2_surj[symmetric] setsum_reindex[OF e2_inj]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   810
    unfolding M_def comp_def apply(intro setsum_cong) apply force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   811
    by (metis e2_surj b1 h1 h2 imageI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   812
    also have "... = N1 b1" using b1 ct1[OF f1b1] by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   813
    finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   814
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   815
    fix b2 assume b2: "b2 \<in> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   816
    hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   817
    by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   818
    have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   819
    unfolding e1_surj[symmetric] setsum_reindex[OF e1_inj]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   820
    unfolding M_def comp_def apply(intro setsum_cong) apply force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   821
    by (metis e1_surj b2 h1 h2 imageI)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   822
    also have "... = N2 b2" using b2 ct2[OF f2b2] by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   823
    finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   824
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   825
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   826
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   827
lemma supp_vimage_mmap:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   828
assumes "M \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   829
shows "supp M \<subseteq> f -` (supp (mmap f M))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   830
using assms by (auto simp: mmap_image)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   831
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   832
lemma mmap_ge_0:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   833
assumes "M \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   834
shows "0 < mmap f M b \<longleftrightarrow> (\<exists>a. 0 < M a \<and> f a = b)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   835
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   836
  have f: "finite {a. f a = b \<and> 0 < M a}" using assms unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   837
  show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   838
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   839
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   840
lemma finite_twosets:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   841
assumes "finite B1" and "finite B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   842
shows "finite {u b1 b2 |b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"  (is "finite ?A")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   843
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   844
  have A: "?A = (\<lambda> b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   845
  show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   846
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   847
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   848
lemma wp_mmap:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   849
fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   850
assumes wp: "wpull A B1 B2 f1 f2 p1 p2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   851
shows
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   852
"wpull {M. M \<in> multiset \<and> supp M \<subseteq> A}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   853
       {N1. N1 \<in> multiset \<and> supp N1 \<subseteq> B1} {N2. N2 \<in> multiset \<and> supp N2 \<subseteq> B2}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   854
       (mmap f1) (mmap f2) (mmap p1) (mmap p2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   855
unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   856
  fix N1 :: "'b1 \<Rightarrow> nat" and N2 :: "'b2 \<Rightarrow> nat"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   857
  assume mmap': "mmap f1 N1 = mmap f2 N2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   858
  and N1[simp]: "N1 \<in> multiset" "supp N1 \<subseteq> B1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   859
  and N2[simp]: "N2 \<in> multiset" "supp N2 \<subseteq> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   860
  have mN1[simp]: "mmap f1 N1 \<in> multiset" using N1 by (auto simp: mmap)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   861
  have mN2[simp]: "mmap f2 N2 \<in> multiset" using N2 by (auto simp: mmap)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   862
  def P \<equiv> "mmap f1 N1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   863
  have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   864
  note P = P1 P2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   865
  have P_mult[simp]: "P \<in> multiset" unfolding P_def using N1 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   866
  have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   867
  have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   868
  have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   869
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   870
  def set1 \<equiv> "\<lambda> c. {b1 \<in> supp N1. f1 b1 = c}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   871
  have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   872
  have fin_set1: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set1 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   873
  using N1(1) unfolding set1_def multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   874
  have set1_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<noteq> {}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   875
  unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   876
  have supp_N1_set1: "supp N1 = (\<Union> c \<in> supp P. set1 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   877
  using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   878
  hence set1_inclN1: "\<And>c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> supp N1" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   879
  hence set1_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> B1" using N1(2) by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   880
  have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   881
  unfolding set1_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   882
  have setsum_set1: "\<And> c. setsum N1 (set1 c) = P c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   883
  unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   884
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   885
  def set2 \<equiv> "\<lambda> c. {b2 \<in> supp N2. f2 b2 = c}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   886
  have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   887
  have fin_set2: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set2 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   888
  using N2(1) unfolding set2_def multiset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   889
  have set2_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<noteq> {}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   890
  unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   891
  have supp_N2_set2: "supp N2 = (\<Union> c \<in> supp P. set2 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   892
  using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   893
  hence set2_inclN2: "\<And>c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> supp N2" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   894
  hence set2_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> B2" using N2(2) by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   895
  have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   896
  unfolding set2_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   897
  have setsum_set2: "\<And> c. setsum N2 (set2 c) = P c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   898
  unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   899
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   900
  have ss: "\<And> c. c \<in> supp P \<Longrightarrow> setsum N1 (set1 c) = setsum N2 (set2 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   901
  unfolding setsum_set1 setsum_set2 ..
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   902
  have "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   903
          \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   904
  using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   905
  by simp (metis set1 set2 set_rev_mp)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   906
  then obtain uu where uu:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   907
  "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   908
     uu c b1b2 \<in> A \<and> p1 (uu c b1b2) = fst b1b2 \<and> p2 (uu c b1b2) = snd b1b2" by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   909
  def u \<equiv> "\<lambda> c b1 b2. uu c (b1,b2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   910
  have u[simp]:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   911
  "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   912
  "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   913
  "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   914
  using uu unfolding u_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   915
  {fix c assume c: "c \<in> supp P"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   916
   have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   917
     fix b1 b1' b2 b2'
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   918
     assume "{b1, b1'} \<subseteq> set1 c" "{b2, b2'} \<subseteq> set2 c" and 0: "u c b1 b2 = u c b1' b2'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   919
     hence "p1 (u c b1 b2) = b1 \<and> p2 (u c b1 b2) = b2 \<and>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   920
            p1 (u c b1' b2') = b1' \<and> p2 (u c b1' b2') = b2'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   921
     using u(2)[OF c] u(3)[OF c] by simp metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   922
     thus "b1 = b1' \<and> b2 = b2'" using 0 by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   923
   qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   924
  } note inj = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   925
  def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   926
  have fin_sset[simp]: "\<And> c. c \<in> supp P \<Longrightarrow> finite (sset c)" unfolding sset_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   927
  using fin_set1 fin_set2 finite_twosets by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   928
  have sset_A: "\<And> c. c \<in> supp P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   929
  {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   930
   then obtain b1 b2 where b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   931
   and a: "a = u c b1 b2" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   932
   have "p1 a \<in> set1 c" and p2a: "p2 a \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   933
   using ac a b1 b2 c u(2) u(3) by simp+
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   934
   hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   935
   unfolding inj2_def by (metis c u(2) u(3))
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   936
  } note u_p12[simp] = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   937
  {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   938
   hence "p1 a \<in> set1 c" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   939
  }note p1[simp] = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   940
  {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   941
   hence "p2 a \<in> set2 c" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   942
  }note p2[simp] = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   943
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   944
  {fix c assume c: "c \<in> supp P"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   945
   hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = N1 b1) \<and>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   946
               (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = N2 b2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   947
   unfolding sset_def
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   948
   using matrix_setsum_finite[OF set1_NE[OF c] fin_set1[OF c]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   949
                                 set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   950
  }
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   951
  then obtain Ms where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   952
  ss1: "\<And> c b1. \<lbrakk>c \<in> supp P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   953
                   setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   954
  ss2: "\<And> c b2. \<lbrakk>c \<in> supp P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   955
                   setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = N2 b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   956
  by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   957
  def SET \<equiv> "\<Union> c \<in> supp P. sset c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   958
  have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   959
  have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   960
  have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   961
  unfolding SET_def sset_def by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   962
  {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   963
   then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   964
   unfolding SET_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   965
   hence "p1 a \<in> set1 c'" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   966
   hence eq: "c = c'" using p1a c c' set1_disj by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   967
   hence "a \<in> sset c" using ac' by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   968
  } note p1_rev = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   969
  {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   970
   then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   971
   unfolding SET_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   972
   hence "p2 a \<in> set2 c'" unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   973
   hence eq: "c = c'" using p2a c c' set2_disj by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   974
   hence "a \<in> sset c" using ac' by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   975
  } note p2_rev = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   976
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   977
  have "\<forall> a \<in> SET. \<exists> c \<in> supp P. a \<in> sset c" unfolding SET_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   978
  then obtain h where h: "\<forall> a \<in> SET. h a \<in> supp P \<and> a \<in> sset (h a)" by metis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   979
  have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   980
                      \<Longrightarrow> h (u c b1 b2) = c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   981
  by (metis h p2 set2 u(3) u_SET)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   982
  have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   983
                      \<Longrightarrow> h (u c b1 b2) = f1 b1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   984
  using h unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   985
  have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   986
                      \<Longrightarrow> h (u c b1 b2) = f2 b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   987
  using h unfolding sset_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   988
  def M \<equiv> "\<lambda> a. if a \<in> SET \<and> p1 a \<in> supp N1 \<and> p2 a \<in> supp N2 then Ms (h a) a else 0"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   989
  have sM: "supp M \<subseteq> SET" "supp M \<subseteq> p1 -` (supp N1)" "supp M \<subseteq> p2 -` (supp N2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   990
  unfolding M_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   991
  show "\<exists>M. (M \<in> multiset \<and> supp M \<subseteq> A) \<and> mmap p1 M = N1 \<and> mmap p2 M = N2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   992
  proof(rule exI[of _ M], safe)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   993
    show "M \<in> multiset"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   994
    unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   995
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   996
    fix a assume "0 < M a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   997
    thus "a \<in> A" unfolding M_def using SET_A by (cases "a \<in> SET") auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   998
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   999
    show "mmap p1 M = N1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1000
    unfolding mmap_def[abs_def] proof(rule ext)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1001
      fix b1
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1002
      let ?K = "{a. p1 a = b1 \<and> 0 < M a}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1003
      show "setsum M ?K = N1 b1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1004
      proof(cases "b1 \<in> supp N1")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1005
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1006
        hence "?K = {}" using sM(2) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1007
        thus ?thesis using False by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1008
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1009
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1010
        def c \<equiv> "f1 b1"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1011
        have c: "c \<in> supp P" and b1: "b1 \<in> set1 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1012
        unfolding set1_def c_def P1 using True by (auto simp: mmap_image)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1013
        have "setsum M ?K = setsum M {a. p1 a = b1 \<and> a \<in> SET}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1014
        apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1015
        also have "... = setsum M ((\<lambda> b2. u c b1 b2) ` (set2 c))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1016
        apply(rule setsum_cong) using c b1 proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1017
          fix a assume p1a: "p1 a \<in> set1 c" and "0 < P c" and "a \<in> SET"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1018
          hence ac: "a \<in> sset c" using p1_rev by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1019
          hence "a = u c (p1 a) (p2 a)" using c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1020
          moreover have "p2 a \<in> set2 c" using ac c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1021
          ultimately show "a \<in> u c (p1 a) ` set2 c" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1022
        next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1023
          fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1024
          hence "u c b1 b2 \<in> SET" using c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1025
        qed auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1026
        also have "... = setsum (\<lambda> b2. M (u c b1 b2)) (set2 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1027
        unfolding comp_def[symmetric] apply(rule setsum_reindex)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1028
        using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1029
        also have "... = N1 b1" unfolding ss1[OF c b1, symmetric]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1030
          apply(rule setsum_cong[OF refl]) unfolding M_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1031
          using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1032
        finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1033
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1034
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1035
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1036
    show "mmap p2 M = N2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1037
    unfolding mmap_def[abs_def] proof(rule ext)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1038
      fix b2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1039
      let ?K = "{a. p2 a = b2 \<and> 0 < M a}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1040
      show "setsum M ?K = N2 b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1041
      proof(cases "b2 \<in> supp N2")
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1042
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1043
        hence "?K = {}" using sM(3) by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1044
        thus ?thesis using False by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1045
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1046
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1047
        def c \<equiv> "f2 b2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1048
        have c: "c \<in> supp P" and b2: "b2 \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1049
        unfolding set2_def c_def P2 using True by (auto simp: mmap_image)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1050
        have "setsum M ?K = setsum M {a. p2 a = b2 \<and> a \<in> SET}"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1051
        apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1052
        also have "... = setsum M ((\<lambda> b1. u c b1 b2) ` (set1 c))"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1053
        apply(rule setsum_cong) using c b2 proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1054
          fix a assume p2a: "p2 a \<in> set2 c" and "0 < P c" and "a \<in> SET"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1055
          hence ac: "a \<in> sset c" using p2_rev by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1056
          hence "a = u c (p1 a) (p2 a)" using c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1057
          moreover have "p1 a \<in> set1 c" using ac c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1058
          ultimately show "a \<in> (\<lambda>b1. u c b1 (p2 a)) ` set1 c" by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1059
        next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1060
          fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1061
          hence "u c b1 b2 \<in> SET" using c by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1062
        qed auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1063
        also have "... = setsum (M o (\<lambda> b1. u c b1 b2)) (set1 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1064
        apply(rule setsum_reindex)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1065
        using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1066
        also have "... = setsum (\<lambda> b1. M (u c b1 b2)) (set1 c)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1067
        unfolding comp_def[symmetric] by simp
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1068
        also have "... = N2 b2" unfolding ss2[OF c b2, symmetric]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1069
          apply(rule setsum_cong[OF refl]) unfolding M_def set2_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1070
          using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2]
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1071
          unfolding set1_def by fastforce
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1072
        finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1073
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1074
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1075
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1076
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1077
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1078
definition multiset_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1079
"multiset_map h = Abs_multiset \<circ> mmap h \<circ> count"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1080
51836
4d6dcd51dd52 renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
blanchet
parents: 51782
diff changeset
  1081
bnf multiset_map [set_of] "\<lambda>_::'a multiset. natLeq" ["{#}"]
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1082
unfolding multiset_map_def
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1083
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1084
  show "Abs_multiset \<circ> mmap id \<circ> count = id" unfolding mmap_id by (auto simp: count_inverse)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1085
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1086
  fix f g
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1087
  show "Abs_multiset \<circ> mmap (g \<circ> f) \<circ> count =
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1088
        Abs_multiset \<circ> mmap g \<circ> count \<circ> (Abs_multiset \<circ> mmap f \<circ> count)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1089
  unfolding comp_def apply(rule ext)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1090
  by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1091
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1092
  fix M f g assume eq: "\<And>a. a \<in> set_of M \<Longrightarrow> f a = g a"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1093
  thus "(Abs_multiset \<circ> mmap f \<circ> count) M = (Abs_multiset \<circ> mmap g \<circ> count) M" apply auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1094
  unfolding cIm_def[abs_def] image_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1095
  by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1096
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1097
  fix f show "set_of \<circ> (Abs_multiset \<circ> mmap f \<circ> count) = op ` f \<circ> set_of"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1098
  by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1099
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1100
  show "card_order natLeq" by (rule natLeq_card_order)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1101
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1102
  show "cinfinite natLeq" by (rule natLeq_cinfinite)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1103
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1104
  fix M show "|set_of M| \<le>o natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1105
  apply(rule ordLess_imp_ordLeq)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1106
  unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1107
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1108
  fix A :: "'a set"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1109
  have "|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_set_of .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1110
  also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1111
  by (rule list_in_bd)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1112
  finally show "|{M. set_of M \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1113
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1114
  fix A B1 B2 f1 f2 p1 p2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1115
  let ?map = "\<lambda> f. Abs_multiset \<circ> mmap f \<circ> count"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1116
  assume wp: "wpull A B1 B2 f1 f2 p1 p2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1117
  show "wpull {x. set_of x \<subseteq> A} {x. set_of x \<subseteq> B1} {x. set_of x \<subseteq> B2}
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1118
              (?map f1) (?map f2) (?map p1) (?map p2)"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1119
  unfolding wpull_def proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1120
    fix y1 y2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1121
    assume y1: "set_of y1 \<subseteq> B1" and y2: "set_of y2 \<subseteq> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1122
    and m: "?map f1 y1 = ?map f2 y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1123
    def N1 \<equiv> "count y1"  def N2 \<equiv> "count y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1124
    have "N1 \<in> multiset \<and> supp N1 \<subseteq> B1" and "N2 \<in> multiset \<and> supp N2 \<subseteq> B2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1125
    and "mmap f1 N1 = mmap f2 N2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1126
    using y1 y2 m unfolding N1_def N2_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1127
    by (auto simp: Abs_multiset_inject count mmap)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1128
    then obtain M where M: "M \<in> multiset \<and> supp M \<subseteq> A"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1129
    and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1130
    using wp_mmap[OF wp] unfolding wpull_def by auto
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1131
    def x \<equiv> "Abs_multiset M"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1132
    show "\<exists>x\<in>{x. set_of x \<subseteq> A}. ?map p1 x = y1 \<and> ?map p2 x = y2"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1133
    apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1134
    by (auto simp: count_inverse Abs_multiset_inverse)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1135
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1136
qed (unfold set_of_empty, auto)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1137
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1138
inductive multiset_rel' where
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1139
Zero: "multiset_rel' R {#} {#}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1140
|
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1141
Plus: "\<lbrakk>R a b; multiset_rel' R M N\<rbrakk> \<Longrightarrow> multiset_rel' R (M + {#a#}) (N + {#b#})"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1142
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1143
lemma multiset_map_Zero_iff[simp]: "multiset_map f M = {#} \<longleftrightarrow> M = {#}"
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
  1144
by (metis image_is_empty multiset.set_map' set_of_eq_empty_iff)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1145
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1146
lemma multiset_map_Zero[simp]: "multiset_map f {#} = {#}" by simp
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1147
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1148
lemma multiset_rel_Zero: "multiset_rel R {#} {#}"
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1149
unfolding multiset_rel_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1150
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1151
declare multiset.count[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1152
declare mmap[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1153
declare Abs_multiset_inverse[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1154
declare multiset.count_inverse[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1155
declare union_preserves_multiset[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1156
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1157
lemma mmap_Plus[simp]:
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1158
assumes "K \<in> multiset" and "L \<in> multiset"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1159
shows "mmap f (\<lambda>a. K a + L a) a = mmap f K a + mmap f L a"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1160
proof-
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1161
  have "{aa. f aa = a \<and> (0 < K aa \<or> 0 < L aa)} \<subseteq>
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1162
        {aa. 0 < K aa} \<union> {aa. 0 < L aa}" (is "?C \<subseteq> ?A \<union> ?B") by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1163
  moreover have "finite (?A \<union> ?B)" apply(rule finite_UnI)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1164
  using assms unfolding multiset_def by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1165
  ultimately have C: "finite ?C" using finite_subset by blast
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1166
  have "setsum K {aa. f aa = a \<and> 0 < K aa} = setsum K {aa. f aa = a \<and> 0 < K aa + L aa}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1167
  apply(rule setsum_mono_zero_cong_left) using C by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1168
  moreover
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1169
  have "setsum L {aa. f aa = a \<and> 0 < L aa} = setsum L {aa. f aa = a \<and> 0 < K aa + L aa}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1170
  apply(rule setsum_mono_zero_cong_left) using C by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1171
  ultimately show ?thesis
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51446
diff changeset
  1172
  unfolding mmap_def by (auto simp add: setsum.distrib)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1173
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1174
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1175
lemma multiset_map_Plus[simp]:
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1176
"multiset_map f (M1 + M2) = multiset_map f M1 + multiset_map f M2"
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1177
unfolding multiset_map_def
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1178
apply(subst multiset.count_inject[symmetric])
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1179
unfolding plus_multiset.rep_eq comp_def by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1180
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1181
lemma multiset_map_singl[simp]: "multiset_map f {#a#} = {#f a#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1182
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1183
  have 0: "\<And> b. card {aa. a = aa \<and> (a = aa \<longrightarrow> f aa = b)} =
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1184
                (if b = f a then 1 else 0)" by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1185
  thus ?thesis
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1186
  unfolding multiset_map_def comp_def mmap_def[abs_def] map_fun_def
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1187
  by (simp, simp add: single_def)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1188
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1189
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1190
lemma multiset_rel_Plus:
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1191
assumes ab: "R a b" and MN: "multiset_rel R M N"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1192
shows "multiset_rel R (M + {#a#}) (N + {#b#})"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1193
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1194
  {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1195
   hence "\<exists>ya. multiset_map fst y + {#a#} = multiset_map fst ya \<and>
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1196
               multiset_map snd y + {#b#} = multiset_map snd ya \<and>
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1197
               set_of ya \<subseteq> {(x, y). R x y}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1198
   apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1199
  }
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1200
  thus ?thesis
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1201
  using assms
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1202
  unfolding multiset_rel_def Grp_def by force
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1203
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1204
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1205
lemma multiset_rel'_imp_multiset_rel:
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1206
"multiset_rel' R M N \<Longrightarrow> multiset_rel R M N"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1207
apply(induct rule: multiset_rel'.induct)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1208
using multiset_rel_Zero multiset_rel_Plus by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1209
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1210
lemma mcard_multiset_map[simp]: "mcard (multiset_map f M) = mcard M"
51548
757fa47af981 centralized various multiset operations in theory multiset;
haftmann
parents: 51489
diff changeset
  1211
proof -
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1212
  def A \<equiv> "\<lambda> b. {a. f a = b \<and> a \<in># M}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1213
  let ?B = "{b. 0 < setsum (count M) (A b)}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1214
  have "{b. \<exists>a. f a = b \<and> a \<in># M} \<subseteq> f ` {a. a \<in># M}" by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1215
  moreover have "finite (f ` {a. a \<in># M})" apply(rule finite_imageI)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1216
  using finite_Collect_mem .
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1217
  ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1218
  have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1219
  by (metis (lifting, mono_tags) mem_Collect_eq rel_simps(54)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1220
                                 setsum_gt_0_iff setsum_infinite)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1221
  have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1222
  apply safe
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1223
    apply (metis less_not_refl setsum_gt_0_iff setsum_infinite)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1224
    by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1225
  hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1226
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1227
  have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = setsum (setsum (count M) o A) ?B"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1228
  unfolding comp_def ..
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1229
  also have "... = (\<Sum>x\<in> A ` ?B. setsum (count M) x)"
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51446
diff changeset
  1230
  unfolding setsum.reindex [OF i, symmetric] ..
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1231
  also have "... = setsum (count M) (\<Union>x\<in>A ` {b. 0 < setsum (count M) (A b)}. x)"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1232
  (is "_ = setsum (count M) ?J")
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51446
diff changeset
  1233
  apply(rule setsum.UNION_disjoint[symmetric])
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1234
  using 0 fin unfolding A_def by (auto intro!: finite_imageI)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1235
  also have "?J = {a. a \<in># M}" unfolding AB unfolding A_def by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1236
  finally have "setsum (\<lambda> x. setsum (count M) (A x)) ?B =
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1237
                setsum (count M) {a. a \<in># M}" .
51548
757fa47af981 centralized various multiset operations in theory multiset;
haftmann
parents: 51489
diff changeset
  1238
  then show ?thesis by (simp add: A_def mcard_unfold_setsum multiset_map_def set_of_def mmap_def)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1239
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1240
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1241
lemma multiset_rel_mcard:
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1242
assumes "multiset_rel R M N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1243
shows "mcard M = mcard N"
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1244
using assms unfolding multiset_rel_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1245
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1246
lemma multiset_induct2[case_names empty addL addR]:
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1247
assumes empty: "P {#} {#}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1248
and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1249
and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1250
shows "P M N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1251
apply(induct N rule: multiset_induct)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1252
  apply(induct M rule: multiset_induct, rule empty, erule addL)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1253
  apply(induct M rule: multiset_induct, erule addR, erule addR)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1254
done
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1255
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1256
lemma multiset_induct2_mcard[consumes 1, case_names empty add]:
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1257
assumes c: "mcard M = mcard N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1258
and empty: "P {#} {#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1259
and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1260
shows "P M N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1261
using c proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1262
  case (less M)  show ?case
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1263
  proof(cases "M = {#}")
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1264
    case True hence "N = {#}" using less.prems by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1265
    thus ?thesis using True empty by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1266
  next
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1267
    case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1268
    have "N \<noteq> {#}" using False less.prems by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1269
    then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1270
    have "mcard M1 = mcard N1" using less.prems unfolding M N by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1271
    thus ?thesis using M N less.hyps add by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1272
  qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1273
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1274
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1275
lemma msed_map_invL:
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1276
assumes "multiset_map f (M + {#a#}) = N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1277
shows "\<exists> N1. N = N1 + {#f a#} \<and> multiset_map f M = N1"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1278
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1279
  have "f a \<in># N"
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
  1280
  using assms multiset.set_map'[of f "M + {#a#}"] by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1281
  then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1282
  have "multiset_map f M = N1" using assms unfolding N by simp
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1283
  thus ?thesis using N by blast
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1284
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1285
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1286
lemma msed_map_invR:
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1287
assumes "multiset_map f M = N + {#b#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1288
shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> multiset_map f M1 = N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1289
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1290
  obtain a where a: "a \<in># M" and fa: "f a = b"
51766
f19a4d0ab1bf renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
blanchet
parents: 51548
diff changeset
  1291
  using multiset.set_map'[of f M] unfolding assms
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1292
  by (metis image_iff mem_set_of_iff union_single_eq_member)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1293
  then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1294
  have "multiset_map f M1 = N" using assms unfolding M fa[symmetric] by simp
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1295
  thus ?thesis using M fa by blast
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1296
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1297
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1298
lemma msed_rel_invL:
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1299
assumes "multiset_rel R (M + {#a#}) N"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1300
shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> multiset_rel R M N1"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1301
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1302
  obtain K where KM: "multiset_map fst K = M + {#a#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1303
  and KN: "multiset_map snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1304
  using assms
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1305
  unfolding multiset_rel_def Grp_def by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1306
  obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1307
  and K1M: "multiset_map fst K1 = M" using msed_map_invR[OF KM] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1308
  obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "multiset_map snd K1 = N1"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1309
  using msed_map_invL[OF KN[unfolded K]] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1310
  have Rab: "R a (snd ab)" using sK a unfolding K by auto
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1311
  have "multiset_rel R M N1" using sK K1M K1N1
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1312
  unfolding K multiset_rel_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1313
  thus ?thesis using N Rab by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1314
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1315
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1316
lemma msed_rel_invR:
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1317
assumes "multiset_rel R M (N + {#b#})"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1318
shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> multiset_rel R M1 N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1319
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1320
  obtain K where KN: "multiset_map snd K = N + {#b#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1321
  and KM: "multiset_map fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1322
  using assms
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1323
  unfolding multiset_rel_def Grp_def by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1324
  obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1325
  and K1N: "multiset_map snd K1 = N" using msed_map_invR[OF KN] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1326
  obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "multiset_map fst K1 = M1"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1327
  using msed_map_invL[OF KM[unfolded K]] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1328
  have Rab: "R (fst ab) b" using sK b unfolding K by auto
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1329
  have "multiset_rel R M1 N" using sK K1N K1M1
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
  1330
  unfolding K multiset_rel_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1331
  thus ?thesis using M Rab by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1332
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1333
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1334
lemma multiset_rel_imp_multiset_rel':
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1335
assumes "multiset_rel R M N"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1336
shows "multiset_rel' R M N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1337
using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1338
  case (less M)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1339
  have c: "mcard M = mcard N" using multiset_rel_mcard[OF less.prems] .
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1340
  show ?case
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1341
  proof(cases "M = {#}")
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1342
    case True hence "N = {#}" using c by simp
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1343
    thus ?thesis using True multiset_rel'.Zero by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1344
  next
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1345
    case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1346
    obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "multiset_rel R M1 N1"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1347
    using msed_rel_invL[OF less.prems[unfolded M]] by auto
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1348
    have "multiset_rel' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1349
    thus ?thesis using multiset_rel'.Plus[of R a b, OF R] unfolding M N by simp
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1350
  qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1351
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1352
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1353
lemma multiset_rel_multiset_rel':
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1354
"multiset_rel R M N = multiset_rel' R M N"
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1355
using  multiset_rel_imp_multiset_rel' multiset_rel'_imp_multiset_rel by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1356
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1357
(* The main end product for multiset_rel: inductive characterization *)
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1358
theorems multiset_rel_induct[case_names empty add, induct pred: multiset_rel] =
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1359
         multiset_rel'.induct[unfolded multiset_rel_multiset_rel'[symmetric]]
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1360
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1361
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1362
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1363
(* Advanced relator customization *)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1364
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1365
(* Set vs. sum relators: *)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1366
(* FIXME: All such facts should be declared as simps: *)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1367
declare sum_rel_simps[simp]
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1368
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1369
lemma set_rel_sum_rel[simp]: 
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1370
"set_rel (sum_rel \<chi> \<phi>) A1 A2 \<longleftrightarrow> 
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1371
 set_rel \<chi> (Inl -` A1) (Inl -` A2) \<and> set_rel \<phi> (Inr -` A1) (Inr -` A2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1372
(is "?L \<longleftrightarrow> ?Rl \<and> ?Rr")
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1373
proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1374
  assume L: "?L"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1375
  show ?Rl unfolding set_rel_def Bex_def vimage_eq proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1376
    fix l1 assume "Inl l1 \<in> A1"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1377
    then obtain a2 where a2: "a2 \<in> A2" and "sum_rel \<chi> \<phi> (Inl l1) a2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1378
    using L unfolding set_rel_def by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1379
    then obtain l2 where "a2 = Inl l2 \<and> \<chi> l1 l2" by (cases a2, auto)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1380
    thus "\<exists> l2. Inl l2 \<in> A2 \<and> \<chi> l1 l2" using a2 by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1381
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1382
    fix l2 assume "Inl l2 \<in> A2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1383
    then obtain a1 where a1: "a1 \<in> A1" and "sum_rel \<chi> \<phi> a1 (Inl l2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1384
    using L unfolding set_rel_def by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1385
    then obtain l1 where "a1 = Inl l1 \<and> \<chi> l1 l2" by (cases a1, auto)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1386
    thus "\<exists> l1. Inl l1 \<in> A1 \<and> \<chi> l1 l2" using a1 by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1387
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1388
  show ?Rr unfolding set_rel_def Bex_def vimage_eq proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1389
    fix r1 assume "Inr r1 \<in> A1"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1390
    then obtain a2 where a2: "a2 \<in> A2" and "sum_rel \<chi> \<phi> (Inr r1) a2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1391
    using L unfolding set_rel_def by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1392
    then obtain r2 where "a2 = Inr r2 \<and> \<phi> r1 r2" by (cases a2, auto)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1393
    thus "\<exists> r2. Inr r2 \<in> A2 \<and> \<phi> r1 r2" using a2 by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1394
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1395
    fix r2 assume "Inr r2 \<in> A2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1396
    then obtain a1 where a1: "a1 \<in> A1" and "sum_rel \<chi> \<phi> a1 (Inr r2)"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1397
    using L unfolding set_rel_def by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1398
    then obtain r1 where "a1 = Inr r1 \<and> \<phi> r1 r2" by (cases a1, auto)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1399
    thus "\<exists> r1. Inr r1 \<in> A1 \<and> \<phi> r1 r2" using a1 by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1400
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1401
next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1402
  assume Rl: "?Rl" and Rr: "?Rr"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1403
  show ?L unfolding set_rel_def Bex_def vimage_eq proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1404
    fix a1 assume a1: "a1 \<in> A1"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1405
    show "\<exists> a2. a2 \<in> A2 \<and> sum_rel \<chi> \<phi> a1 a2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1406
    proof(cases a1)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1407
      case (Inl l1) then obtain l2 where "Inl l2 \<in> A2 \<and> \<chi> l1 l2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1408
      using Rl a1 unfolding set_rel_def by blast
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1409
      thus ?thesis unfolding Inl by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1410
    next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1411
      case (Inr r1) then obtain r2 where "Inr r2 \<in> A2 \<and> \<phi> r1 r2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1412
      using Rr a1 unfolding set_rel_def by blast
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1413
      thus ?thesis unfolding Inr by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1414
    qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1415
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1416
    fix a2 assume a2: "a2 \<in> A2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1417
    show "\<exists> a1. a1 \<in> A1 \<and> sum_rel \<chi> \<phi> a1 a2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1418
    proof(cases a2)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1419
      case (Inl l2) then obtain l1 where "Inl l1 \<in> A1 \<and> \<chi> l1 l2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1420
      using Rl a2 unfolding set_rel_def by blast
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1421
      thus ?thesis unfolding Inl by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1422
    next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1423
      case (Inr r2) then obtain r1 where "Inr r1 \<in> A1 \<and> \<phi> r1 r2"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1424
      using Rr a2 unfolding set_rel_def by blast
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1425
      thus ?thesis unfolding Inr by auto
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1426
    qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1427
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1428
qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1429
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1430
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1431
end