src/HOL/BNF/More_BNFs.thy
author blanchet
Tue, 19 Nov 2013 14:11:26 +0100
changeset 54490 930409d43211
parent 54489 03ff4d1e6784
child 54539 bbab2ebda234
permissions -rw-r--r--
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
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
53124
9ae9bbaccee1 tuned theory imports
traytel
parents: 53026
diff changeset
    15
  Basic_BNFs
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
    16
  "~~/src/HOL/Library/FSet"
49310
6e30078de4f0 renamed "Ordinals_and_Cardinals" to "Cardinals"
blanchet
parents: 49309
diff changeset
    17
  "~~/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
    18
  Countable_Type
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    19
begin
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    20
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    21
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
    22
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
    23
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    24
bnf "'a option"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    25
  map: Option.map
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    26
  sets: Option.set
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    27
  bd: natLeq 
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    28
  wits: None
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
    29
  rel: option_rel
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    30
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    31
  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
    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
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    34
  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
    35
    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
    36
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    37
  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
    38
  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
    39
  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
    40
    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
    41
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    42
  fix f
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    43
  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
    44
    by fastforce
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
  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
    47
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    48
  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
    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 x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    51
  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
    52
    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
    53
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    54
  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
    55
  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
    56
  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
    57
    (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
    58
    (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
    59
    unfolding wpull_def
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    60
  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
    61
    fix b1 b2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    62
    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
    63
    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
    64
      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
    65
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    66
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    67
  fix z
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    68
  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
    69
  thus False by simp
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
    70
next
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
    71
  fix R
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    72
  show "option_rel R =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
    73
        (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
    74
         Grp {x. Option.set x \<subseteq> Collect (split R)} (Option.map snd)"
53026
e1a548c11845 got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
traytel
parents: 53013
diff changeset
    75
  unfolding option_rel_def 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
    76
  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
    77
           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
    78
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
    79
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    80
lemma wpull_map:
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    81
  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
    82
  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
    83
    (is "wpull ?A ?B1 ?B2 _ _ _ _")
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    84
proof (unfold wpull_def)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    85
  { 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
    86
    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
    87
    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
    88
    proof (induct as bs rule: list_induct2)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    89
      case (Cons a as b bs)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    90
      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
    91
      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
    92
      moreover
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    93
      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
    94
      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
    95
      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
    96
    qed simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    97
  }
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
    98
  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
    99
    (\<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
   100
qed
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   101
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   102
bnf "'a list"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   103
  map: map
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   104
  sets: set
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   105
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   106
  wits: Nil
54424
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   107
  rel: list_all2
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   108
proof -
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   109
  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
   110
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   111
  fix f g
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   112
  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
   113
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   114
  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
   115
  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
   116
  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
   117
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   118
  fix f
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   119
  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
   120
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   121
  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
   122
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   123
  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
   124
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   125
  fix x
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   126
  show "|set x| \<le>o natLeq"
52660
7f7311d04727 killed unused theorems
traytel
parents: 52659
diff changeset
   127
    by (metis List.finite_set finite_iff_ordLess_natLeq ordLess_imp_ordLeq)
54424
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   128
next
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   129
  fix R
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   130
  show "list_all2 R =
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   131
         (Grp {x. set x \<subseteq> {(x, y). R x y}} (map fst))\<inverse>\<inverse> OO
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   132
         Grp {x. set x \<subseteq> {(x, y). R x y}} (map snd)"
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   133
    unfolding list_all2_def[abs_def] Grp_def fun_eq_iff relcompp.simps conversep.simps
72ec50a85f3b standard relator for list bnf
traytel
parents: 54421
diff changeset
   134
    by (force simp: zip_map_fst_snd)
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   135
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
   136
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   137
(* Finite sets *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   138
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   139
lemma wpull_image:
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   140
  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
   141
  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
   142
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
   143
  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
   144
  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
   145
  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
   146
  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
   147
    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
   148
    proof
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   149
      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
   150
      proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   151
        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
   152
        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
   153
        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
   154
        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
   155
        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
   156
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   157
    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
   158
    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
   159
    proof
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   160
      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
   161
      proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   162
        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
   163
        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
   164
        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
   165
        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
   166
        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
   167
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   168
    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
   169
  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
   170
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   171
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   172
context
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   173
includes fset.lifting
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   174
begin
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   175
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   176
lemma fset_rel_alt: "fset_rel R a b \<longleftrightarrow> (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and>
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   177
                                        (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   178
  by transfer (simp add: set_rel_def)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   179
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   180
lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   181
  apply (rule f_the_inv_into_f[unfolded inj_on_def])
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   182
  apply (simp add: fset_inject) apply (rule range_eqI Abs_fset_inverse[symmetric] CollectI)+
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   183
  .
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   184
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   185
lemma fset_rel_aux:
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   186
"(\<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>
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   187
 ((Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage fst))\<inverse>\<inverse> OO
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   188
  Grp {a. fset a \<subseteq> {(a, b). R a b}} (fimage snd)) a b" (is "?L = ?R")
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   189
proof
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   190
  assume ?L
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   191
  def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'")
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   192
  have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) (transfer, simp)+
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   193
  hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   194
  show ?R unfolding Grp_def relcompp.simps conversep.simps
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   195
  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   196
    from * show "a = fimage fst R'" using conjunct1[OF `?L`]
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   197
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   198
    from * show "b = fimage snd R'" using conjunct2[OF `?L`]
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   199
      by (transfer, auto simp add: image_def Int_def split: prod.splits)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   200
  qed (auto simp add: *)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   201
next
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   202
  assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   203
  apply (simp add: subset_eq Ball_def)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   204
  apply (rule conjI)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   205
  apply (transfer, clarsimp, metis snd_conv)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   206
  by (transfer, clarsimp, metis fst_conv)
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   207
qed
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   208
54430
e9ff6a25aaa6 renamed thm
blanchet
parents: 54424
diff changeset
   209
lemma wpull_fimage:
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   210
  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
   211
  shows "wpull {x. fset x \<subseteq> A} {x. fset x \<subseteq> B1} {x. fset x \<subseteq> B2}
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   212
              (fimage f1) (fimage f2) (fimage p1) (fimage p2)"
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   213
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
   214
  fix y1 y2
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   215
  assume Y1: "fset y1 \<subseteq> B1" and Y2: "fset y2 \<subseteq> B2"
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   216
  assume "fimage f1 y1 = fimage f2 y2"
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   217
  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
   218
  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
   219
    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
   220
    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
   221
  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
   222
  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
   223
  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
   224
  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
   225
  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
   226
  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
   227
  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
   228
  have fX': "finite X'" unfolding X'_def by transfer simp
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   229
  then obtain x where X'eq: "X' = fset x" by transfer simp
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   230
  show "\<exists>x. fset x \<subseteq> A \<and> fimage p1 x = y1 \<and> fimage p2 x = y2"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   231
     using X' Y1 Y2 by (auto simp: X'eq intro!: exI[of _ "x"]) (transfer, blast)+
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   232
qed
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   233
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   234
bnf "'a fset"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   235
  map: fimage
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   236
  sets: fset 
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   237
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   238
  wits: "{||}"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   239
  rel: fset_rel
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   240
apply -
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   241
          apply transfer' apply simp
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   242
         apply transfer' apply force
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   243
        apply transfer apply force
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   244
       apply transfer' apply force
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   245
      apply (rule natLeq_card_order)
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   246
     apply (rule natLeq_cinfinite)
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   247
    apply transfer apply (metis ordLess_imp_ordLeq finite_iff_ordLess_natLeq)
54430
e9ff6a25aaa6 renamed thm
blanchet
parents: 54424
diff changeset
   248
  apply (erule wpull_fimage)
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   249
 apply (simp add: Grp_def relcompp.simps conversep.simps fun_eq_iff fset_rel_alt fset_rel_aux) 
51410
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   250
apply transfer apply simp
f0865a641e76 BNF uses fset defined via Lifting/Transfer rather than Quotient
traytel
parents: 51371
diff changeset
   251
done
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   252
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   253
lemma fset_rel_fset: "set_rel \<chi> (fset A1) (fset A2) = fset_rel \<chi> A1 A2"
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   254
  by transfer (rule refl)
51371
197ad6b8f763 some simp rules for fset
traytel
parents: 50144
diff changeset
   255
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   256
end
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   257
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53374
diff changeset
   258
lemmas [simp] = fset.map_comp fset.map_id fset.set_map
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
   259
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   260
(* Countable sets *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   261
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   262
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
   263
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
   264
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
   265
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
   266
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
   267
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   268
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
   269
"|{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
   270
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
   271
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
   272
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
   273
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   274
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
   275
"|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
   276
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
   277
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   278
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
   279
"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
   280
apply default
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   281
 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
   282
 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
   283
 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
   284
 apply assumption
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   285
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
   286
apply (rule disjI1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   287
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
   288
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   289
lemma rcset_to_rcset: "countable A \<Longrightarrow> rcset (the_inv rcset A) = A"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   290
  apply (rule f_the_inv_into_f[unfolded inj_on_def image_iff])
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   291
   apply transfer' apply simp
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   292
  apply transfer' apply simp
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   293
  done
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   294
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   295
lemma Collect_Int_Times:
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   296
"{(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
   297
by auto
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   298
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
   299
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
   300
"cset_rel R a b \<longleftrightarrow>
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   301
 (\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and>
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   302
 (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t)"
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   303
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
   304
lemma cset_rel_aux:
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   305
"(\<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>
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   306
 ((Grp {x. rcset x \<subseteq> {(a, b). R a b}} (cimage fst))\<inverse>\<inverse> OO
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   307
          Grp {x. rcset x \<subseteq> {(a, b). R a b}} (cimage snd)) a b" (is "?L = ?R")
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   308
proof
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   309
  assume ?L
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   310
  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
   311
  (is "the_inv rcset ?L'")
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   312
  have L: "countable ?L'" by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   313
  hence *: "rcset R' = ?L'" unfolding R'_def using fset_to_fset by (intro rcset_to_rcset)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   314
  thus ?R unfolding Grp_def relcompp.simps conversep.simps
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   315
  proof (intro CollectI prod_caseI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   316
    from * `?L` show "a = cimage fst R'" by transfer (auto simp: image_def Collect_Int_Times)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   317
  next
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   318
    from * `?L` show "b = cimage snd R'" by transfer (auto simp: image_def Collect_Int_Times)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   319
  qed simp_all
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   320
next
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   321
  assume ?R thus ?L unfolding Grp_def relcompp.simps conversep.simps
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   322
    by transfer force
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   323
qed
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   324
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   325
bnf "'a cset"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   326
  map: cimage
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   327
  sets: rcset
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   328
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   329
  wits: "cempty"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   330
  rel: cset_rel
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   331
proof -
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   332
  show "cimage id = id" by transfer' simp
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   333
next
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   334
  fix f g show "cimage (g \<circ> f) = cimage g \<circ> cimage f" by transfer' fastforce
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   335
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   336
  fix C f g assume eq: "\<And>a. a \<in> rcset C \<Longrightarrow> f a = g a"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   337
  thus "cimage f C = cimage g C" by transfer force
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   338
next
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   339
  fix f show "rcset \<circ> cimage f = op ` f \<circ> rcset" by transfer' fastforce
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   340
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   341
  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
   342
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   343
  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
   344
next
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   345
  fix C show "|rcset C| \<le>o natLeq" by transfer (unfold 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
   346
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   347
  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
   348
  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
   349
  show "wpull {x. rcset x \<subseteq> A} {x. rcset x \<subseteq> B1} {x. rcset x \<subseteq> B2}
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   350
              (cimage f1) (cimage f2) (cimage p1) (cimage p2)"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   351
  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
   352
    fix y1 y2
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   353
    assume Y1: "rcset y1 \<subseteq> B1" and Y2: "rcset y2 \<subseteq> B2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   354
    assume "cimage f1 y1 = cimage f2 y2"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   355
    hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)" by transfer
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   356
    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
   357
    and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   358
    using wpull_image[OF wp] unfolding wpull_def Pow_def Bex_def mem_Collect_eq
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   359
      by (auto elim!: allE[of _ "rcset y1"] allE[of _ "rcset y2"])
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   360
    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
   361
    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
   362
    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
   363
    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
   364
    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
   365
    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
   366
    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
   367
    have fX': "countable X'" unfolding X'_def by simp
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   368
    then obtain x where X'eq: "X' = rcset x" by transfer blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   369
    show "\<exists>x\<in>{x. rcset x \<subseteq> A}. cimage p1 x = y1 \<and> cimage p2 x = y2"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   370
      using X' Y1 Y2 unfolding X'eq by (intro bexI[of _ "x"]) (transfer, auto)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   371
  qed
49461
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   372
next
de07eecb2664 adapting "More_BNFs" to new relators/predicators
blanchet
parents: 49440
diff changeset
   373
  fix R
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   374
  show "cset_rel R =
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   375
        (Grp {x. rcset x \<subseteq> Collect (split R)} (cimage fst))\<inverse>\<inverse> OO
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   376
         Grp {x. rcset x \<subseteq> Collect (split R)} (cimage snd)"
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51836
diff changeset
   377
  unfolding cset_rel_def[abs_def] cset_rel_aux by simp
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   378
qed (transfer, simp)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   379
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   380
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   381
(* Multisets *)
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 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
   384
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
   385
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
   386
(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
   387
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   388
  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
   389
  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
   390
  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
   391
  finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   392
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   393
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   394
lift_definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" is
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   395
  "\<lambda>h f b. setsum f {a. h a = b \<and> f a > 0} :: nat"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   396
unfolding multiset_def proof safe
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   397
  fix h :: "'a \<Rightarrow> 'b" and f :: "'a \<Rightarrow> nat"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   398
  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
   399
  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
   400
  (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
   401
  proof- let ?B = "{b. 0 < setsum f (?As b)}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   402
    have "\<And> b. finite (?As b)" using fin by simp
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49878
diff changeset
   403
    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
   404
    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
   405
    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
   406
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   407
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   408
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53124
diff changeset
   409
lemma mmap_id0: "mmap id = id"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   410
proof (intro ext multiset_eqI)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   411
  fix f a show "count (mmap id f) a = count (id f) a"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   412
  proof (cases "count f a = 0")
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   413
    case False
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   414
    hence 1: "{aa. aa = a \<and> aa \<in># f} = {a}" by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   415
    thus ?thesis by transfer auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   416
  qed (transfer, simp)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   417
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   418
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   419
lemma inj_on_setsum_inv:
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   420
assumes 1: "(0::nat) < setsum (count f) {a. h a = b' \<and> a \<in># f}" (is "0 < setsum (count f) ?A'")
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   421
and     2: "{a. h a = b \<and> a \<in># f} = {a. h a = b' \<and> a \<in># f}" (is "?A = ?A'")
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   422
shows "b = b'"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   423
using assms 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
   424
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   425
lemma mmap_comp:
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   426
fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   427
shows "mmap (h2 o h1) = mmap h2 o mmap h1"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   428
proof (intro ext multiset_eqI)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   429
  fix f :: "'a multiset" fix c :: 'c
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   430
  let ?A = "{a. h2 (h1 a) = c \<and> a \<in># f}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   431
  let ?As = "\<lambda> b. {a. h1 a = b \<and> a \<in># f}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   432
  let ?B = "{b. h2 b = c \<and> 0 < setsum (count f) (?As b)}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   433
  have 0: "{?As b | b.  b \<in> ?B} = ?As ` ?B" by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   434
  have "\<And> b. finite (?As b)" by transfer (simp add: multiset_def)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   435
  hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   436
  hence A: "?A = \<Union> {?As b | b.  b \<in> ?B}" by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   437
  have "setsum (count f) ?A = setsum (setsum (count f)) {?As b | b.  b \<in> ?B}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   438
    unfolding A by transfer (intro setsum_Union_disjoint, auto simp: multiset_def)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   439
  also have "... = setsum (setsum (count f)) (?As ` ?B)" unfolding 0 ..
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   440
  also have "... = setsum (setsum (count f) o ?As) ?B"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   441
    by(intro setsum_reindex) (auto simp add: setsum_gt_0_iff inj_on_def)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   442
  also have "... = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" unfolding comp_def ..
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   443
  finally have "setsum (count f) ?A = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" .
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   444
  thus "count (mmap (h2 \<circ> h1) f) c = count ((mmap h2 \<circ> mmap h1) f) c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   445
    by transfer (unfold o_apply, blast)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   446
qed
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   447
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   448
lemma mmap_cong:
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   449
assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   450
shows "mmap f M = mmap g M"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   451
using assms by transfer (auto intro!: setsum_cong)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   452
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   453
context
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   454
begin
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   455
interpretation lifting_syntax .
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   456
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   457
lemma set_of_transfer[transfer_rule]: "(pcr_multiset op = ===> op =) (\<lambda>f. {a. 0 < f a}) set_of"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   458
  unfolding set_of_def pcr_multiset_def cr_multiset_def fun_rel_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   459
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   460
end
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 52662
diff changeset
   461
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   462
lemma set_of_mmap: "set_of o mmap h = image h o set_of"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   463
proof (rule ext, unfold o_apply)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   464
  fix M show "set_of (mmap h M) = h ` set_of M"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   465
    by transfer (auto simp add: multiset_def setsum_gt_0_iff)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   466
qed
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   467
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   468
lemma multiset_of_surj:
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   469
  "multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   470
proof safe
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   471
  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
   472
  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
   473
  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
   474
  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
   475
next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   476
  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
   477
  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
   478
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   479
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   480
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
   481
"|{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
   482
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
   483
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   484
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
   485
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
   486
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
   487
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   488
  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
   489
  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
   490
  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
   491
  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
   492
  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
   493
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   494
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   495
lemma matrix_count:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   496
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
   497
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
   498
shows
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   499
"\<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
   500
       (\<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
   501
(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
   502
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   503
  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
   504
        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
   505
  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
   506
"\<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
   507
     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
   508
      clarify)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   509
  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
   510
  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
   511
                \<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
   512
                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
   513
  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
   514
  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
   515
  proof(cases n1)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   516
    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
   517
    show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   518
    proof(cases n2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   519
      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
   520
      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
   521
      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
   522
    next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   523
      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
   524
      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
   525
      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
   526
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   527
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   528
    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
   529
    show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   530
    proof(cases n2)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   531
      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
   532
      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
   533
      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
   534
    next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   535
      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
   536
      show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   537
      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
   538
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   539
        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
   540
        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
   541
        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
   542
        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
   543
        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
   544
        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
   545
        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
   546
        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
   547
                                       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
   548
        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
   549
        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
   550
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   551
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   552
        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
   553
        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
   554
        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
   555
        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
   556
        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
   557
        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
   558
        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
   559
        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
   560
        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
   561
                                       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
   562
        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
   563
        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
   564
      qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   565
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   566
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   567
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   568
  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
   569
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   570
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   571
definition
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   572
"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
   573
 \<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
   574
                  \<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
   575
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   576
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
   577
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
   578
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
   579
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
   580
            (\<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
   581
            (\<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
   582
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   583
  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
   584
  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
   585
  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
   586
  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
   587
  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
   588
  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
   589
  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
   590
  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
   591
  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
   592
  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
   593
  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
   594
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   595
  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
   596
  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
   597
  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
   598
  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
   599
  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
   600
  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
   601
  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
   602
  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
   603
  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
   604
  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
   605
  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
   606
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   607
  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
   608
  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
   609
  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
   610
  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
   611
  obtain ct where
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   612
  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
   613
  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
   614
  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
   615
  (*  *)
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   616
  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
   617
  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
   618
  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
   619
  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
   620
  "\<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
   621
  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
   622
  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
   623
                  "\<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
   624
  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
   625
  {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
   626
   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
   627
   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
   628
   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
   629
   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
   630
   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
   631
  }
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   632
  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
   633
        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
   634
  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
   635
  show ?thesis
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   636
  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
   637
    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
   638
    hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
53124
9ae9bbaccee1 tuned theory imports
traytel
parents: 53026
diff changeset
   639
    by (metis image_eqI lessThan_iff less_Suc_eq_le)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   640
    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
   641
    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
   642
    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
   643
    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
   644
    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
   645
    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
   646
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   647
    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
   648
    hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
53124
9ae9bbaccee1 tuned theory imports
traytel
parents: 53026
diff changeset
   649
    by (metis image_eqI lessThan_iff less_Suc_eq_le)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   650
    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
   651
    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
   652
    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
   653
    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
   654
    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
   655
    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
   656
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   657
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   658
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   659
lemma supp_vimage_mmap: "set_of M \<subseteq> f -` (set_of (mmap f M))"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   660
  by transfer (auto simp: multiset_def 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
   661
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   662
lemma mmap_ge_0: "b \<in># mmap f M \<longleftrightarrow> (\<exists>a. a \<in># M \<and> f a = b)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   663
  by transfer (auto simp: multiset_def 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
   664
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   665
lemma finite_twosets:
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   666
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
   667
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
   668
proof-
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   669
  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
   670
  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
   671
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   672
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   673
lemma wpull_mmap:
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   674
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
   675
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
   676
shows
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   677
"wpull {M. set_of M \<subseteq> A}
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   678
       {N1. set_of N1 \<subseteq> B1} {N2. set_of N2 \<subseteq> B2}
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   679
       (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
   680
unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   681
  fix N1 :: "'b1 multiset" and N2 :: "'b2 multiset"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   682
  assume mmap': "mmap f1 N1 = mmap f2 N2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   683
  and N1[simp]: "set_of N1 \<subseteq> B1"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   684
  and N2[simp]: "set_of N2 \<subseteq> B2"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   685
  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
   686
  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
   687
  note P = P1 P2
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   688
  have fin_N1[simp]: "finite (set_of N1)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   689
   and fin_N2[simp]: "finite (set_of N2)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   690
   and fin_P[simp]: "finite (set_of P)" by auto
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   691
  (*  *)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   692
  def set1 \<equiv> "\<lambda> c. {b1 \<in> set_of N1. f1 b1 = c}"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   693
  have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   694
  have fin_set1: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (set1 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   695
    using N1(1) unfolding set1_def multiset_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   696
  have set1_NE: "\<And> c. c \<in> set_of P \<Longrightarrow> set1 c \<noteq> {}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   697
   unfolding set1_def set_of_def P mmap_ge_0 by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   698
  have supp_N1_set1: "set_of N1 = (\<Union> c \<in> set_of P. set1 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   699
    using supp_vimage_mmap[of N1 f1] unfolding set1_def P1 by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   700
  hence set1_inclN1: "\<And>c. c \<in> set_of P \<Longrightarrow> set1 c \<subseteq> set_of N1" by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   701
  hence set1_incl: "\<And> c. c \<in> set_of P \<Longrightarrow> set1 c \<subseteq> B1" using N1 by blast
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   702
  have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   703
    unfolding set1_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   704
  have setsum_set1: "\<And> c. setsum (count N1) (set1 c) = count P c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   705
    unfolding P1 set1_def by transfer (auto intro: setsum_cong)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   706
  (*  *)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   707
  def set2 \<equiv> "\<lambda> c. {b2 \<in> set_of N2. f2 b2 = c}"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   708
  have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   709
  have fin_set2: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (set2 c)"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   710
  using N2(1) unfolding set2_def multiset_def by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   711
  have set2_NE: "\<And> c. c \<in> set_of P \<Longrightarrow> set2 c \<noteq> {}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   712
    unfolding set2_def P2 mmap_ge_0 set_of_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   713
  have supp_N2_set2: "set_of N2 = (\<Union> c \<in> set_of P. set2 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   714
    using supp_vimage_mmap[of N2 f2] unfolding set2_def P2 by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   715
  hence set2_inclN2: "\<And>c. c \<in> set_of P \<Longrightarrow> set2 c \<subseteq> set_of N2" by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   716
  hence set2_incl: "\<And> c. c \<in> set_of P \<Longrightarrow> set2 c \<subseteq> B2" using N2 by blast
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   717
  have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   718
    unfolding set2_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   719
  have setsum_set2: "\<And> c. setsum (count N2) (set2 c) = count P c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   720
    unfolding P2 set2_def by transfer (auto intro: setsum_cong)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   721
  (*  *)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   722
  have ss: "\<And> c. c \<in> set_of P \<Longrightarrow> setsum (count N1) (set1 c) = setsum (count N2) (set2 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   723
    unfolding setsum_set1 setsum_set2 ..
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   724
  have "\<forall> c \<in> set_of P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   725
          \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   726
    using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   727
    by simp (metis set1 set2 set_rev_mp)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   728
  then obtain uu where uu:
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   729
  "\<forall> c \<in> set_of P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   730
     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
   731
  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
   732
  have u[simp]:
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   733
  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   734
  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   735
  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   736
    using uu unfolding u_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   737
  {fix c assume c: "c \<in> set_of P"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   738
   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
   739
     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
   740
     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
   741
     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
   742
            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
   743
     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
   744
     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
   745
   qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   746
  } note inj = this
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   747
  def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   748
  have fin_sset[simp]: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (sset c)" unfolding sset_def
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   749
    using fin_set1 fin_set2 finite_twosets by blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   750
  have sset_A: "\<And> c. c \<in> set_of P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   751
  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   752
   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
   753
   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
   754
   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
   755
   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
   756
   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
   757
   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
   758
  } note u_p12[simp] = this
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   759
  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   760
   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
   761
  }note p1[simp] = this
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   762
  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   763
   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
   764
  }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
   765
  (*  *)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   766
  {fix c assume c: "c \<in> set_of P"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   767
   hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = count N1 b1) \<and>
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   768
               (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = count N2 b2)"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   769
   unfolding sset_def
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   770
   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
   771
                                 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
   772
  }
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   773
  then obtain Ms where
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   774
  ss1: "\<And> c b1. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow>
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   775
                   setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = count N1 b1" and
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   776
  ss2: "\<And> c b2. \<lbrakk>c \<in> set_of P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow>
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   777
                   setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = count N2 b2"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   778
  by metis
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   779
  def SET \<equiv> "\<Union> c \<in> set_of P. sset c"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   780
  have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   781
  have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   782
  have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   783
    unfolding SET_def sset_def by blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   784
  {fix c a assume c: "c \<in> set_of P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   785
   then obtain c' where c': "c' \<in> set_of P" and ac': "a \<in> sset c'"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   786
    unfolding SET_def by auto
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   787
   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
   788
   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
   789
   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
   790
  } note p1_rev = this
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   791
  {fix c a assume c: "c \<in> set_of P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   792
   then obtain c' where c': "c' \<in> set_of P" and ac': "a \<in> sset c'"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   793
   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
   794
   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
   795
   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
   796
   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
   797
  } 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
   798
  (*  *)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   799
  have "\<forall> a \<in> SET. \<exists> c \<in> set_of P. a \<in> sset c" unfolding SET_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   800
  then obtain h where h: "\<forall> a \<in> SET. h a \<in> set_of P \<and> a \<in> sset (h a)" by metis
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   801
  have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   802
                      \<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
   803
  by (metis h p2 set2 u(3) u_SET)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   804
  have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   805
                      \<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
   806
  using h unfolding sset_def by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   807
  have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   808
                      \<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
   809
  using h unfolding sset_def by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   810
  def M \<equiv>
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   811
    "Abs_multiset (\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   812
  have "(\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0) \<in> multiset"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   813
    unfolding multiset_def by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   814
  hence [transfer_rule]: "pcr_multiset op = (\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0) M"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   815
    unfolding M_def pcr_multiset_def cr_multiset_def by (auto simp: Abs_multiset_inverse)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   816
  have sM: "set_of M \<subseteq> SET" "set_of M \<subseteq> p1 -` (set_of N1)" "set_of M \<subseteq> p2 -` set_of N2"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   817
    by (transfer, auto split: split_if_asm)+
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   818
  show "\<exists>M. set_of M \<subseteq> A \<and> mmap p1 M = N1 \<and> mmap p2 M = N2"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   819
  proof(rule exI[of _ M], safe)
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   820
    fix a assume *: "a \<in> set_of M"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   821
    from SET_A show "a \<in> A"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   822
    proof (cases "a \<in> SET")
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   823
      case False thus ?thesis using * by transfer' auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   824
    qed blast
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   825
  next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   826
    show "mmap p1 M = N1"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   827
    proof(intro multiset_eqI)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   828
      fix b1
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   829
      let ?K = "{a. p1 a = b1 \<and> a \<in># M}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   830
      have "setsum (count M) ?K = count N1 b1"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   831
      proof(cases "b1 \<in> set_of N1")
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   832
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   833
        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
   834
        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
   835
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   836
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   837
        def c \<equiv> "f1 b1"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   838
        have c: "c \<in> set_of P" and b1: "b1 \<in> set1 c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   839
          unfolding set1_def c_def P1 using True by (auto simp: o_eq_dest[OF set_of_mmap])
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   840
        with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p1 a = b1 \<and> a \<in> SET}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   841
          by transfer (force intro: setsum_mono_zero_cong_left split: split_if_asm)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   842
        also have "... = setsum (count M) ((\<lambda> b2. u c b1 b2) ` (set2 c))"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   843
          apply(rule setsum_cong) using c b1 proof safe
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   844
          fix a assume p1a: "p1 a \<in> set1 c" and "c \<in> set_of P" and "a \<in> SET"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   845
          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
   846
          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
   847
          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
   848
          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
   849
        qed auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   850
        also have "... = setsum (\<lambda> b2. count M (u c b1 b2)) (set2 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   851
          unfolding comp_def[symmetric] apply(rule setsum_reindex)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   852
          using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   853
        also have "... = count N1 b1" unfolding ss1[OF c b1, symmetric]
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   854
          apply(rule setsum_cong[OF refl]) apply (transfer fixing: Ms u c b1 set2)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   855
          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
   856
        finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   857
      qed
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   858
      thus "count (mmap p1 M) b1 = count N1 b1" by transfer
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   859
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   860
  next
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   861
next
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   862
    show "mmap p2 M = N2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   863
    proof(intro multiset_eqI)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   864
      fix b2
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   865
      let ?K = "{a. p2 a = b2 \<and> a \<in># M}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   866
      have "setsum (count M) ?K = count N2 b2"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   867
      proof(cases "b2 \<in> set_of N2")
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   868
        case False
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   869
        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
   870
        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
   871
      next
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   872
        case True
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   873
        def c \<equiv> "f2 b2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   874
        have c: "c \<in> set_of P" and b2: "b2 \<in> set2 c"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   875
          unfolding set2_def c_def P2 using True by (auto simp: o_eq_dest[OF set_of_mmap])
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   876
        with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p2 a = b2 \<and> a \<in> SET}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   877
          by transfer (force intro: setsum_mono_zero_cong_left split: split_if_asm)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   878
        also have "... = setsum (count M) ((\<lambda> b1. u c b1 b2) ` (set1 c))"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   879
          apply(rule setsum_cong) using c b2 proof safe
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   880
          fix a assume p2a: "p2 a \<in> set2 c" and "c \<in> set_of P" and "a \<in> SET"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   881
          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
   882
          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
   883
          moreover have "p1 a \<in> set1 c" using ac c by auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   884
          ultimately show "a \<in> (\<lambda>x. u c x (p2 a)) ` set1 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
   885
        qed auto
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   886
        also have "... = setsum (count M o (\<lambda> b1. u c b1 b2)) (set1 c)"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   887
          apply(rule setsum_reindex)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   888
          using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   889
        also have "... = setsum (\<lambda> b1. count M (u c b1 b2)) (set1 c)" by simp
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   890
        also have "... = count N2 b2" unfolding ss2[OF c b2, symmetric] o_def
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   891
          apply(rule setsum_cong[OF refl]) apply (transfer fixing: Ms u c b2 set1)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   892
          using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] set1_def by fastforce
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   893
        finally show ?thesis .
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   894
      qed
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   895
      thus "count (mmap p2 M) b2 = count N2 b2" by transfer
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   896
    qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   897
  qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   898
qed
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   899
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   900
lemma set_of_bd: "|set_of x| \<le>o natLeq"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   901
  by transfer
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   902
    (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   903
54421
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   904
bnf "'a multiset"
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   905
  map: mmap
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   906
  sets: set_of 
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   907
  bd: natLeq
632be352a5a3 more explicit syntax for defining a bnf
traytel
parents: 54014
diff changeset
   908
  wits: "{#}"
53270
c8628119d18e renamed BNF axiom
blanchet
parents: 53124
diff changeset
   909
by (auto simp add: mmap_id0 mmap_comp set_of_mmap natLeq_card_order natLeq_cinfinite set_of_bd
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   910
  intro: mmap_cong wpull_mmap)
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
   911
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   912
inductive rel_multiset' where
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   913
Zero: "rel_multiset' R {#} {#}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   914
|
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   915
Plus: "\<lbrakk>R a b; rel_multiset' R M N\<rbrakk> \<Longrightarrow> rel_multiset' R (M + {#a#}) (N + {#b#})"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   916
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   917
lemma map_multiset_Zero_iff[simp]: "mmap f M = {#} \<longleftrightarrow> M = {#}"
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53288
diff changeset
   918
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
   919
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   920
lemma map_multiset_Zero[simp]: "mmap f {#} = {#}" by simp
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   921
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   922
lemma rel_multiset_Zero: "rel_multiset R {#} {#}"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   923
unfolding rel_multiset_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   924
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   925
declare multiset.count[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   926
declare Abs_multiset_inverse[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   927
declare multiset.count_inverse[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   928
declare union_preserves_multiset[simp]
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   929
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   930
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   931
lemma map_multiset_Plus[simp]: "mmap f (M1 + M2) = mmap f M1 + mmap f M2"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   932
proof (intro multiset_eqI, transfer fixing: f)
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   933
  fix x :: 'a and M1 M2 :: "'b \<Rightarrow> nat"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   934
  assume "M1 \<in> multiset" "M2 \<in> multiset"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   935
  hence "setsum M1 {a. f a = x \<and> 0 < M1 a} = setsum M1 {a. f a = x \<and> 0 < M1 a + M2 a}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   936
        "setsum M2 {a. f a = x \<and> 0 < M2 a} = setsum M2 {a. f a = x \<and> 0 < M1 a + M2 a}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   937
    by (auto simp: multiset_def intro!: setsum_mono_zero_cong_left)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53290
diff changeset
   938
  then show "(\<Sum>a | f a = x \<and> 0 < M1 a + M2 a. M1 a + M2 a) =
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   939
       setsum M1 {a. f a = x \<and> 0 < M1 a} +
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   940
       setsum M2 {a. f a = x \<and> 0 < M2 a}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   941
    by (auto simp: setsum.distrib[symmetric])
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   942
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   943
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   944
lemma map_multiset_singl[simp]: "mmap f {#a#} = {#f a#}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   945
  by transfer auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   946
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   947
lemma rel_multiset_Plus:
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   948
assumes ab: "R a b" and MN: "rel_multiset R M N"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   949
shows "rel_multiset R (M + {#a#}) (N + {#b#})"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   950
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   951
  {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   952
   hence "\<exists>ya. mmap fst y + {#a#} = mmap fst ya \<and>
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   953
               mmap snd y + {#b#} = mmap snd ya \<and>
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   954
               set_of ya \<subseteq> {(x, y). R x y}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   955
   apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   956
  }
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   957
  thus ?thesis
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
   958
  using assms
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   959
  unfolding rel_multiset_def Grp_def by force
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   960
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   961
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   962
lemma rel_multiset'_imp_rel_multiset:
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   963
"rel_multiset' R M N \<Longrightarrow> rel_multiset R M N"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   964
apply(induct rule: rel_multiset'.induct)
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   965
using rel_multiset_Zero rel_multiset_Plus by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   966
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   967
lemma mcard_mmap[simp]: "mcard (mmap f M) = mcard M"
51548
757fa47af981 centralized various multiset operations in theory multiset;
haftmann
parents: 51489
diff changeset
   968
proof -
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   969
  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
   970
  let ?B = "{b. 0 < setsum (count M) (A b)}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   971
  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
   972
  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
   973
  using finite_Collect_mem .
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   974
  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
   975
  have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54430
diff changeset
   976
    by (metis (lifting, full_types) mem_Collect_eq neq0_conv setsum.neutral)
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   977
  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
   978
  apply safe
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   979
    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
   980
    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
   981
  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
   982
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   983
  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
   984
  unfolding comp_def ..
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   985
  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
   986
  unfolding setsum.reindex [OF i, symmetric] ..
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   987
  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
   988
  (is "_ = setsum (count M) ?J")
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51446
diff changeset
   989
  apply(rule setsum.UNION_disjoint[symmetric])
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   990
  using 0 fin unfolding A_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   991
  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
   992
  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
   993
                setsum (count M) {a. a \<in># M}" .
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
   994
  then show ?thesis unfolding mcard_unfold_setsum A_def by transfer
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   995
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   996
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   997
lemma rel_multiset_mcard:
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
   998
assumes "rel_multiset R M N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
   999
shows "mcard M = mcard N"
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1000
using assms unfolding rel_multiset_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1001
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1002
lemma multiset_induct2[case_names empty addL addR]:
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
  1003
assumes empty: "P {#} {#}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1004
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
  1005
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
  1006
shows "P M N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1007
apply(induct N rule: multiset_induct)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1008
  apply(induct M rule: multiset_induct, rule empty, erule addL)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1009
  apply(induct M rule: multiset_induct, erule addR, erule addR)
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1010
done
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1011
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1012
lemma multiset_induct2_mcard[consumes 1, case_names empty add]:
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1013
assumes c: "mcard M = mcard N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1014
and empty: "P {#} {#}"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1015
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
  1016
shows "P M N"
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1017
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
  1018
  case (less M)  show ?case
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1019
  proof(cases "M = {#}")
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1020
    case True hence "N = {#}" using less.prems by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1021
    thus ?thesis using True empty by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1022
  next
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1023
    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
  1024
    have "N \<noteq> {#}" using False less.prems by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1025
    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
  1026
    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
  1027
    thus ?thesis using M N less.hyps add by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1028
  qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1029
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1030
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1031
lemma msed_map_invL:
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1032
assumes "mmap f (M + {#a#}) = N"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1033
shows "\<exists> N1. N = N1 + {#f a#} \<and> mmap f M = N1"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1034
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1035
  have "f a \<in># N"
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53288
diff changeset
  1036
  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
  1037
  then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1038
  have "mmap f M = N1" using assms unfolding N by simp
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1039
  thus ?thesis using N by blast
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1040
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1041
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1042
lemma msed_map_invR:
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1043
assumes "mmap f M = N + {#b#}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1044
shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> mmap f M1 = N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1045
proof-
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1046
  obtain a where a: "a \<in># M" and fa: "f a = b"
53290
b6c3be868217 renamed BNF fact
blanchet
parents: 53288
diff changeset
  1047
  using multiset.set_map[of f M] unfolding assms
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1048
  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
  1049
  then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1050
  have "mmap f M1 = N" using assms unfolding M fa[symmetric] by simp
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1051
  thus ?thesis using M fa by blast
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1052
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1053
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1054
lemma msed_rel_invL:
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1055
assumes "rel_multiset R (M + {#a#}) N"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1056
shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_multiset R M N1"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1057
proof-
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1058
  obtain K where KM: "mmap fst K = M + {#a#}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1059
  and KN: "mmap snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1060
  using assms
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1061
  unfolding rel_multiset_def Grp_def by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1062
  obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1063
  and K1M: "mmap fst K1 = M" using msed_map_invR[OF KM] by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1064
  obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "mmap snd K1 = N1"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1065
  using msed_map_invL[OF KN[unfolded K]] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1066
  have Rab: "R a (snd ab)" using sK a unfolding K by auto
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1067
  have "rel_multiset R M N1" using sK K1M K1N1
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1068
  unfolding K rel_multiset_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1069
  thus ?thesis using N Rab by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1070
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1071
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1072
lemma msed_rel_invR:
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1073
assumes "rel_multiset R M (N + {#b#})"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1074
shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_multiset R M1 N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1075
proof-
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1076
  obtain K where KN: "mmap snd K = N + {#b#}"
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1077
  and KM: "mmap fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1078
  using assms
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1079
  unfolding rel_multiset_def Grp_def by auto
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1080
  obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
52662
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1081
  and K1N: "mmap snd K1 = N" using msed_map_invR[OF KN] by auto
c7cae5ce217d use transfer/lifting for proving countable set and multisets being BNFs
traytel
parents: 52660
diff changeset
  1082
  obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "mmap fst K1 = M1"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1083
  using msed_map_invL[OF KM[unfolded K]] by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1084
  have Rab: "R (fst ab) b" using sK b unfolding K by auto
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1085
  have "rel_multiset R M1 N" using sK K1N K1M1
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1086
  unfolding K rel_multiset_def Grp_def by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1087
  thus ?thesis using M Rab by auto
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1088
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1089
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1090
lemma rel_multiset_imp_rel_multiset':
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1091
assumes "rel_multiset R M N"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1092
shows "rel_multiset' R M N"
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1093
using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
49463
83ac281bcdc2 provide predicator, define relator
blanchet
parents: 49461
diff changeset
  1094
  case (less M)
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1095
  have c: "mcard M = mcard N" using rel_multiset_mcard[OF less.prems] .
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1096
  show ?case
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1097
  proof(cases "M = {#}")
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1098
    case True hence "N = {#}" using c by simp
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1099
    thus ?thesis using True rel_multiset'.Zero by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1100
  next
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1101
    case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1102
    obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_multiset R M1 N1"
49507
8826d5a4332b renamed "pred" to "rel" (relator)
blanchet
parents: 49463
diff changeset
  1103
    using msed_rel_invL[OF less.prems[unfolded M]] by auto
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1104
    have "rel_multiset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1105
    thus ?thesis using rel_multiset'.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
  1106
  qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1107
qed
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1108
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1109
lemma rel_multiset_rel_multiset':
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1110
"rel_multiset R M N = rel_multiset' R M N"
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1111
using  rel_multiset_imp_rel_multiset' rel_multiset'_imp_rel_multiset by auto
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1112
54490
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1113
(* The main end product for rel_multiset: inductive characterization *)
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1114
theorems rel_multiset_induct[case_names empty add, induct pred: rel_multiset] =
930409d43211 use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
blanchet
parents: 54489
diff changeset
  1115
         rel_multiset'.induct[unfolded rel_multiset_rel_multiset'[symmetric]]
49440
4ff2976f4056 Added missing predicators (for multisets and countable sets)
popescua
parents: 49434
diff changeset
  1116
49877
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1117
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1118
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1119
(* Advanced relator customization *)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1120
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1121
(* Set vs. sum relators: *)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1122
(* 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
  1123
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
  1124
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1125
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
  1126
"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
  1127
 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
  1128
(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
  1129
proof safe
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1130
  assume L: "?L"
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1131
  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
  1132
    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
  1133
    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
  1134
    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
  1135
    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
  1136
    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
  1137
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1138
    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
  1139
    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
  1140
    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
  1141
    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
  1142
    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
  1143
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1144
  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
  1145
    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
  1146
    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
  1147
    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
  1148
    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
  1149
    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
  1150
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1151
    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
  1152
    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
  1153
    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
  1154
    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
  1155
    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
  1156
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1157
next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1158
  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
  1159
  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
  1160
    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
  1161
    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
  1162
    proof(cases a1)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1163
      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
  1164
      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
  1165
      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
  1166
    next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1167
      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
  1168
      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
  1169
      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
  1170
    qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1171
  next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1172
    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
  1173
    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
  1174
    proof(cases a2)
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1175
      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
  1176
      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
  1177
      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
  1178
    next
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1179
      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
  1180
      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
  1181
      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
  1182
    qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1183
  qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1184
qed
b75555ec30a4 ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
popescua
parents: 49514
diff changeset
  1185
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
diff changeset
  1186
end