src/ZF/Induct/Multiset.thy
author wenzelm
Mon, 09 Oct 2006 02:19:51 +0200
changeset 20898 113c9516a2d7
parent 18415 eb68dc98bda2
child 24892 c663e675e177
permissions -rw-r--r--
attribute symmetric: zero_var_indexes;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12089
diff changeset
     1
(*  Title:      ZF/Induct/Multiset.thy
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     2
    ID:         $Id$
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     4
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     5
A definitional theory of multisets,
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     6
including a wellfoundedness proof for the multiset order.
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     7
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     8
The theory features ordinal multisets and the usual ordering.
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     9
*)
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    10
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    11
theory Multiset
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    12
imports FoldSet Acc
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    13
begin
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    14
12860
7fc3fbfed8ef Multiset: added the translation Mult(A) => A-||>nat-{0}
paulson
parents: 12610
diff changeset
    15
consts
7fc3fbfed8ef Multiset: added the translation Mult(A) => A-||>nat-{0}
paulson
parents: 12610
diff changeset
    16
  (* Short cut for multiset space *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    17
  Mult :: "i=>i"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    18
translations
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12891
diff changeset
    19
  "Mult(A)" => "A -||> nat-{0}"
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    20
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    21
constdefs
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    22
12891
92af5c3a10fb a new definition of "restrict"
paulson
parents: 12860
diff changeset
    23
  (* This is the original "restrict" from ZF.thy.
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    24
     Restricts the function f to the domain A
12891
92af5c3a10fb a new definition of "restrict"
paulson
parents: 12860
diff changeset
    25
     FIXME: adapt Multiset to the new "restrict". *)
92af5c3a10fb a new definition of "restrict"
paulson
parents: 12860
diff changeset
    26
92af5c3a10fb a new definition of "restrict"
paulson
parents: 12860
diff changeset
    27
  funrestrict :: "[i,i] => i"
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    28
  "funrestrict(f,A) == \<lambda>x \<in> A. f`x"
12891
92af5c3a10fb a new definition of "restrict"
paulson
parents: 12860
diff changeset
    29
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    30
  (* M is a multiset *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    31
  multiset :: "i => o"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    32
  "multiset(M) == \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    33
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    34
  mset_of :: "i=>i"
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    35
  "mset_of(M) == domain(M)"
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    36
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    37
  munion    :: "[i, i] => i" (infixl "+#" 65)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    38
  "M +# N == \<lambda>x \<in> mset_of(M) Un mset_of(N).
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    39
     if x \<in> mset_of(M) Int mset_of(N) then  (M`x) #+ (N`x)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    40
     else (if x \<in> mset_of(M) then M`x else N`x)"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    41
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12891
diff changeset
    42
  (*convert a function to a multiset by eliminating 0*)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    43
  normalize :: "i => i"
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12891
diff changeset
    44
  "normalize(f) ==
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    45
       if (\<exists>A. f \<in> A -> nat & Finite(A)) then
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    46
            funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12891
diff changeset
    47
       else 0"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    48
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    49
  mdiff  :: "[i, i] => i" (infixl "-#" 65)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    50
  "M -# N ==  normalize(\<lambda>x \<in> mset_of(M).
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    51
			if x \<in> mset_of(N) then M`x #- N`x else M`x)"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    52
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    53
  (* set of elements of a multiset *)
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    54
  msingle :: "i => i"    ("{#_#}")
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    55
  "{#a#} == {<a, 1>}"
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    56
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    57
  MCollect :: "[i, i=>o] => i"  (*comprehension*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    58
  "MCollect(M, P) == funrestrict(M, {x \<in> mset_of(M). P(x)})"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    59
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    60
  (* Counts the number of occurences of an element in a multiset *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    61
  mcount :: "[i, i] => i"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    62
  "mcount(M, a) == if a \<in> mset_of(M) then  M`a else 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    63
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    64
  msize :: "i => i"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    65
  "msize(M) == setsum(%a. $# mcount(M,a), mset_of(M))"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    66
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    67
syntax
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    68
  melem :: "[i,i] => o"    ("(_/ :# _)" [50, 51] 50)
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    69
  "@MColl" :: "[pttrn, i, o] => i" ("(1{# _ : _./ _#})")
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    70
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    71
syntax (xsymbols)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    72
  "@MColl" :: "[pttrn, i, o] => i" ("(1{# _ \<in> _./ _#})")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    73
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    74
translations
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    75
  "a :# M" == "a \<in> mset_of(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    76
  "{#x \<in> M. P#}" == "MCollect(M, %x. P)"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    77
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    78
  (* multiset orderings *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    79
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    80
constdefs
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    81
   (* multirel1 has to be a set (not a predicate) so that we can form
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    82
      its transitive closure and reason about wf(.) and acc(.) *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    83
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    84
  multirel1 :: "[i,i]=>i"
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    85
  "multirel1(A, r) ==
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    86
     {<M, N> \<in> Mult(A)*Mult(A).
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    87
      \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    88
      N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    89
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    90
  multirel :: "[i, i] => i"
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    91
  "multirel(A, r) == multirel1(A, r)^+" 		
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    92
12860
7fc3fbfed8ef Multiset: added the translation Mult(A) => A-||>nat-{0}
paulson
parents: 12610
diff changeset
    93
  (* ordinal multiset orderings *)
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    94
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    95
  omultiset :: "i => o"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    96
  "omultiset(M) == \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    97
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    98
  mless :: "[i, i] => o" (infixl "<#" 50)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
    99
  "M <# N ==  \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   100
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   101
  mle  :: "[i, i] => o"  (infixl "<#=" 50)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   102
  "M <#= N == (omultiset(M) & M = N) | M <# N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   103
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   104
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   105
subsection{*Properties of the original "restrict" from ZF.thy*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   106
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   107
lemma funrestrict_subset: "[| f \<in> Pi(C,B);  A\<subseteq>C |] ==> funrestrict(f,A) \<subseteq> f"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   108
by (auto simp add: funrestrict_def lam_def intro: apply_Pair)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   109
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   110
lemma funrestrict_type:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   111
    "[| !!x. x \<in> A ==> f`x \<in> B(x) |] ==> funrestrict(f,A) \<in> Pi(A,B)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   112
by (simp add: funrestrict_def lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   113
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   114
lemma funrestrict_type2: "[| f \<in> Pi(C,B);  A\<subseteq>C |] ==> funrestrict(f,A) \<in> Pi(A,B)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   115
by (blast intro: apply_type funrestrict_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   116
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   117
lemma funrestrict [simp]: "a \<in> A ==> funrestrict(f,A) ` a = f`a"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   118
by (simp add: funrestrict_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   119
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   120
lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   121
by (simp add: funrestrict_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   122
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   123
lemma domain_funrestrict [simp]: "domain(funrestrict(f,C)) = C"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   124
by (auto simp add: funrestrict_def lam_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   125
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   126
lemma fun_cons_funrestrict_eq:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   127
     "f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, funrestrict(f, b))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   128
apply (rule equalityI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   129
prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   130
apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   131
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   132
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   133
declare domain_of_fun [simp]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   134
declare domainE [rule del]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   135
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   136
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   137
text{* A useful simplification rule *}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   138
lemma multiset_fun_iff:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   139
     "(f \<in> A -> nat-{0}) <-> f \<in> A->nat&(\<forall>a \<in> A. f`a \<in> nat & 0 < f`a)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   140
apply safe
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   141
apply (rule_tac [4] B1 = "range (f) " in Pi_mono [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   142
apply (auto intro!: Ord_0_lt
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   143
            dest: apply_type Diff_subset [THEN Pi_mono, THEN subsetD]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   144
            simp add: range_of_fun apply_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   145
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   146
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   147
(** The multiset space  **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   148
lemma multiset_into_Mult: "[| multiset(M); mset_of(M)\<subseteq>A |] ==> M \<in> Mult(A)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   149
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   150
apply (auto simp add: multiset_fun_iff mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   151
apply (rule_tac B1 = "nat-{0}" in FiniteFun_mono [THEN subsetD], simp_all)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   152
apply (rule Finite_into_Fin [THEN [2] Fin_mono [THEN subsetD], THEN fun_FiniteFunI])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   153
apply (simp_all (no_asm_simp) add: multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   154
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   155
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   156
lemma Mult_into_multiset: "M \<in> Mult(A) ==> multiset(M) & mset_of(M)\<subseteq>A"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   157
apply (simp add: multiset_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   158
apply (frule FiniteFun_is_fun)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   159
apply (drule FiniteFun_domain_Fin)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   160
apply (frule FinD, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   161
apply (rule_tac x = "domain (M) " in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   162
apply (blast intro: Fin_into_Finite)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   163
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   164
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   165
lemma Mult_iff_multiset: "M \<in> Mult(A) <-> multiset(M) & mset_of(M)\<subseteq>A"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   166
by (blast dest: Mult_into_multiset intro: multiset_into_Mult)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   167
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   168
lemma multiset_iff_Mult_mset_of: "multiset(M) <-> M \<in> Mult(mset_of(M))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   169
by (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   170
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   171
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   172
text{*The @{term multiset} operator*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   173
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   174
(* the empty multiset is 0 *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   175
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   176
lemma multiset_0 [simp]: "multiset(0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   177
by (auto intro: FiniteFun.intros simp add: multiset_iff_Mult_mset_of)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   178
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   179
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   180
text{*The @{term mset_of} operator*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   181
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   182
lemma multiset_set_of_Finite [simp]: "multiset(M) ==> Finite(mset_of(M))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   183
by (simp add: multiset_def mset_of_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   184
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   185
lemma mset_of_0 [iff]: "mset_of(0) = 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   186
by (simp add: mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   187
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   188
lemma mset_is_0_iff: "multiset(M) ==> mset_of(M)=0 <-> M=0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   189
by (auto simp add: multiset_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   190
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   191
lemma mset_of_single [iff]: "mset_of({#a#}) = {a}"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   192
by (simp add: msingle_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   193
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   194
lemma mset_of_union [iff]: "mset_of(M +# N) = mset_of(M) Un mset_of(N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   195
by (simp add: mset_of_def munion_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   196
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   197
lemma mset_of_diff [simp]: "mset_of(M)\<subseteq>A ==> mset_of(M -# N) \<subseteq> A"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   198
by (auto simp add: mdiff_def multiset_def normalize_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   199
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   200
(* msingle *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   202
lemma msingle_not_0 [iff]: "{#a#} \<noteq> 0 & 0 \<noteq> {#a#}"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   203
by (simp add: msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   204
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   205
lemma msingle_eq_iff [iff]: "({#a#} = {#b#}) <->  (a = b)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   206
by (simp add: msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   207
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   208
lemma msingle_multiset [iff,TC]: "multiset({#a#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   209
apply (simp add: multiset_def msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   210
apply (rule_tac x = "{a}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   211
apply (auto intro: Finite_cons Finite_0 fun_extend3)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   212
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   213
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   214
(** normalize **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   215
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   216
lemmas Collect_Finite = Collect_subset [THEN subset_Finite, standard]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   217
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   218
lemma normalize_idem [simp]: "normalize(normalize(f)) = normalize(f)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   219
apply (simp add: normalize_def funrestrict_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   220
apply (case_tac "\<exists>A. f \<in> A -> nat & Finite (A) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   221
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   222
apply (drule_tac x = "{x \<in> domain (f) . 0 < f ` x}" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   223
apply auto
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   224
apply (auto  intro!: lam_type simp add: Collect_Finite)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   225
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   226
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   227
lemma normalize_multiset [simp]: "multiset(M) ==> normalize(M) = M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   228
by (auto simp add: multiset_def normalize_def mset_of_def funrestrict_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   229
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   230
lemma multiset_normalize [simp]: "multiset(normalize(f))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   231
apply (simp add: normalize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   232
apply (simp add: normalize_def mset_of_def multiset_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   233
apply (rule_tac x = "{x \<in> A . 0<f`x}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   234
apply (auto intro: Collect_subset [THEN subset_Finite] funrestrict_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   235
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   236
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   237
(** Typechecking rules for union and difference of multisets **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   238
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   239
(* union *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   240
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   241
lemma munion_multiset [simp]: "[| multiset(M); multiset(N) |] ==> multiset(M +# N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   242
apply (unfold multiset_def munion_def mset_of_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   243
apply (rule_tac x = "A Un Aa" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   244
apply (auto intro!: lam_type intro: Finite_Un simp add: multiset_fun_iff zero_less_add)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   245
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   246
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   247
(* difference *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   248
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   249
lemma mdiff_multiset [simp]: "multiset(M -# N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   250
by (simp add: mdiff_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   251
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   252
(** Algebraic properties of multisets **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   253
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   254
(* Union *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   255
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   256
lemma munion_0 [simp]: "multiset(M) ==> M +# 0 = M & 0 +# M = M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   257
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   258
apply (auto simp add: munion_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   259
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   260
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   261
lemma munion_commute: "M +# N = N +# M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   262
by (auto intro!: lam_cong simp add: munion_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   263
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   264
lemma munion_assoc: "(M +# N) +# K = M +# (N +# K)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   265
apply (unfold munion_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   266
apply (rule lam_cong, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   267
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   268
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   269
lemma munion_lcommute: "M +# (N +# K) = N +# (M +# K)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   270
apply (unfold munion_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   271
apply (rule lam_cong, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   272
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   273
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   274
lemmas munion_ac = munion_commute munion_assoc munion_lcommute
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   275
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   276
(* Difference *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   277
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   278
lemma mdiff_self_eq_0 [simp]: "M -# M = 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   279
by (simp add: mdiff_def normalize_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   280
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   281
lemma mdiff_0 [simp]: "0 -# M = 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   282
by (simp add: mdiff_def normalize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   283
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   284
lemma mdiff_0_right [simp]: "multiset(M) ==> M -# 0 = M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   285
by (auto simp add: multiset_def mdiff_def normalize_def multiset_fun_iff mset_of_def funrestrict_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   286
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   287
lemma mdiff_union_inverse2 [simp]: "multiset(M) ==> M +# {#a#} -# {#a#} = M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   288
apply (unfold multiset_def munion_def mdiff_def msingle_def normalize_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   289
apply (auto cong add: if_cong simp add: ltD multiset_fun_iff funrestrict_def subset_Un_iff2 [THEN iffD1])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   290
prefer 2 apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   291
apply (subgoal_tac [2] "{x \<in> A \<union> {a} . x \<noteq> a \<and> x \<in> A} = A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   292
apply (rule fun_extension, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   293
apply (drule_tac x = "A Un {a}" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   294
apply (simp add: Finite_Un)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   295
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   296
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   297
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   298
(** Count of elements **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   299
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   300
lemma mcount_type [simp,TC]: "multiset(M) ==> mcount(M, a) \<in> nat"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   301
by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   302
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   303
lemma mcount_0 [simp]: "mcount(0, a) = 0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   304
by (simp add: mcount_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   305
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   306
lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   307
by (simp add: mcount_def mset_of_def msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   308
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   309
lemma mcount_union [simp]: "[| multiset(M); multiset(N) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   310
                     ==>  mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   311
apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   312
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   313
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   314
lemma mcount_diff [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   315
     "multiset(M) ==> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   316
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   317
apply (auto dest!: not_lt_imp_le
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   318
     simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   319
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   320
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   321
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   322
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   323
lemma mcount_elem: "[| multiset(M); a \<in> mset_of(M) |] ==> 0 < mcount(M, a)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   324
apply (simp add: multiset_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   325
apply (simp add: mcount_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   326
apply (simp add: multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   327
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   328
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   329
(** msize **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   330
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   331
lemma msize_0 [simp]: "msize(0) = #0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   332
by (simp add: msize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   333
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   334
lemma msize_single [simp]: "msize({#a#}) = #1"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   335
by (simp add: msize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   336
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   337
lemma msize_type [simp,TC]: "msize(M) \<in> int"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   338
by (simp add: msize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   339
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   340
lemma msize_zpositive: "multiset(M)==> #0 $\<le> msize(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   341
by (auto simp add: msize_def intro: g_zpos_imp_setsum_zpos)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   342
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   343
lemma msize_int_of_nat: "multiset(M) ==> \<exists>n \<in> nat. msize(M)= $# n"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   344
apply (rule not_zneg_int_of)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   345
apply (simp_all (no_asm_simp) add: msize_type [THEN znegative_iff_zless_0] not_zless_iff_zle msize_zpositive)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   346
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   347
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   348
lemma not_empty_multiset_imp_exist:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   349
     "[| M\<noteq>0; multiset(M) |] ==> \<exists>a \<in> mset_of(M). 0 < mcount(M, a)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   350
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   351
apply (erule not_emptyE)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   352
apply (auto simp add: mset_of_def mcount_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   353
apply (blast dest!: fun_is_rel)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   354
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   355
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   356
lemma msize_eq_0_iff: "multiset(M) ==> msize(M)=#0 <-> M=0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   357
apply (simp add: msize_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   358
apply (rule_tac Pa = "setsum (?u,?v) \<noteq> #0" in swap)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   359
apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   360
apply (drule not_empty_multiset_imp_exist, assumption, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   361
apply (subgoal_tac "Finite (mset_of (M) - {a}) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   362
 prefer 2 apply (simp add: Finite_Diff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   363
apply (subgoal_tac "setsum (%x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   364
 prefer 2 apply (simp add: cons_Diff, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   365
apply (subgoal_tac "#0 $\<le> setsum (%x. $# mcount (M, x), mset_of (M) - {a}) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   366
apply (rule_tac [2] g_zpos_imp_setsum_zpos)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   367
apply (auto simp add: Finite_Diff not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   368
apply (rule not_zneg_int_of [THEN bexE])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   369
apply (auto simp del: int_of_0 simp add: int_of_add [symmetric] int_of_0 [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   370
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   371
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   372
lemma setsum_mcount_Int:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   373
     "Finite(A) ==> setsum(%a. $# mcount(N, a), A Int mset_of(N))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   374
		  = setsum(%a. $# mcount(N, a), A)"
18415
eb68dc98bda2 improved proofs;
wenzelm
parents: 16973
diff changeset
   375
apply (induct rule: Finite_induct)
eb68dc98bda2 improved proofs;
wenzelm
parents: 16973
diff changeset
   376
 apply auto
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   377
apply (subgoal_tac "Finite (B Int mset_of (N))")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   378
prefer 2 apply (blast intro: subset_Finite)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   379
apply (auto simp add: mcount_def Int_cons_left)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   380
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   381
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   382
lemma msize_union [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   383
     "[| multiset(M); multiset(N) |] ==> msize(M +# N) = msize(M) $+ msize(N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   384
apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   385
apply (subst Int_commute)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   386
apply (simp add: setsum_mcount_Int)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   387
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   388
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   389
lemma msize_eq_succ_imp_elem: "[|msize(M)= $# succ(n); n \<in> nat|] ==> \<exists>a. a \<in> mset_of(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   390
apply (unfold msize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   391
apply (blast dest: setsum_succD)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   392
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   393
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   394
(** Equality of multisets **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   395
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   396
lemma equality_lemma:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   397
     "[| multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   398
      ==> mset_of(M)=mset_of(N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   399
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   400
apply (rule sym, rule equalityI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   401
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   402
apply (drule_tac [!] x=x in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   403
apply (case_tac [2] "x \<in> Aa", case_tac "x \<in> A", auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   404
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   405
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   406
lemma multiset_equality:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   407
  "[| multiset(M); multiset(N) |]==> M=N<->(\<forall>a. mcount(M, a)=mcount(N, a))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   408
apply auto
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   409
apply (subgoal_tac "mset_of (M) = mset_of (N) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   410
prefer 2 apply (blast intro: equality_lemma)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   411
apply (simp add: multiset_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   412
apply (auto simp add: multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   413
apply (rule fun_extension)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   414
apply (blast, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   415
apply (drule_tac x = x in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   416
apply (auto simp add: mcount_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   417
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   418
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   419
(** More algebraic properties of multisets **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   420
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   421
lemma munion_eq_0_iff [simp]: "[|multiset(M); multiset(N)|]==>(M +# N =0) <-> (M=0 & N=0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   422
by (auto simp add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   423
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   424
lemma empty_eq_munion_iff [simp]: "[|multiset(M); multiset(N)|]==>(0=M +# N) <-> (M=0 & N=0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   425
apply (rule iffI, drule sym)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   426
apply (simp_all add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   427
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   428
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   429
lemma munion_right_cancel [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   430
     "[| multiset(M); multiset(N); multiset(K) |]==>(M +# K = N +# K)<->(M=N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   431
by (auto simp add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   432
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   433
lemma munion_left_cancel [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   434
  "[|multiset(K); multiset(M); multiset(N)|] ==>(K +# M = K +# N) <-> (M = N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   435
by (auto simp add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   436
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   437
lemma nat_add_eq_1_cases: "[| m \<in> nat; n \<in> nat |] ==> (m #+ n = 1) <-> (m=1 & n=0) | (m=0 & n=1)"
18415
eb68dc98bda2 improved proofs;
wenzelm
parents: 16973
diff changeset
   438
by (induct_tac n) auto
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   439
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   440
lemma munion_is_single:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   441
     "[|multiset(M); multiset(N)|] 
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   442
      ==> (M +# N = {#a#}) <->  (M={#a#} & N=0) | (M = 0 & N = {#a#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   443
apply (simp (no_asm_simp) add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   444
apply safe
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   445
apply simp_all
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   446
apply (case_tac "aa=a")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   447
apply (drule_tac [2] x = aa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   448
apply (drule_tac x = a in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   449
apply (simp add: nat_add_eq_1_cases, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   450
apply (case_tac "aaa=aa", simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   451
apply (drule_tac x = aa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   452
apply (simp add: nat_add_eq_1_cases)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   453
apply (case_tac "aaa=a")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   454
apply (drule_tac [4] x = aa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   455
apply (drule_tac [3] x = a in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   456
apply (drule_tac [2] x = aaa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   457
apply (drule_tac x = aa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   458
apply (simp_all add: nat_add_eq_1_cases)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   459
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   460
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   461
lemma msingle_is_union: "[| multiset(M); multiset(N) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   462
  ==> ({#a#} = M +# N) <-> ({#a#} = M  & N=0 | M = 0 & {#a#} = N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   463
apply (subgoal_tac " ({#a#} = M +# N) <-> (M +# N = {#a#}) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   464
apply (simp (no_asm_simp) add: munion_is_single)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   465
apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   466
apply (blast dest: sym)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   467
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   468
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   469
(** Towards induction over multisets **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   470
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   471
lemma setsum_decr:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   472
"Finite(A)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   473
  ==>  (\<forall>M. multiset(M) -->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   474
  (\<forall>a \<in> mset_of(M). setsum(%z. $# mcount(M(a:=M`a #- 1), z), A) =
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   475
  (if a \<in> A then setsum(%z. $# mcount(M, z), A) $- #1
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   476
           else setsum(%z. $# mcount(M, z), A))))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   477
apply (unfold multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   478
apply (erule Finite_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   479
apply (auto simp add: multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   480
apply (unfold mset_of_def mcount_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   481
apply (case_tac "x \<in> A", auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   482
apply (subgoal_tac "$# M ` x $+ #-1 = $# M ` x $- $# 1")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   483
apply (erule ssubst)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   484
apply (rule int_of_diff, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   485
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   486
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   487
lemma setsum_decr2:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   488
     "Finite(A)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   489
      ==> \<forall>M. multiset(M) --> (\<forall>a \<in> mset_of(M).
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15481
diff changeset
   490
           setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) =
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15481
diff changeset
   491
           (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15481
diff changeset
   492
            else setsum(%x. $# mcount(M, x), A)))"
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   493
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   494
apply (erule Finite_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   495
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   496
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   497
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   498
lemma setsum_decr3: "[| Finite(A); multiset(M); a \<in> mset_of(M) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   499
      ==> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   500
          (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   501
           else setsum(%x. $# mcount(M, x), A))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   502
apply (subgoal_tac "setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   503
apply (rule_tac [2] setsum_Diff [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   504
apply (rule sym, rule ssubst, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   505
apply (rule sym, drule setsum_decr2, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   506
apply (simp add: mcount_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   507
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   508
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   509
lemma nat_le_1_cases: "n \<in> nat ==> n le 1 <-> (n=0 | n=1)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   510
by (auto elim: natE)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   511
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   512
lemma succ_pred_eq_self: "[| 0<n; n \<in> nat |] ==> succ(n #- 1) = n"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   513
apply (subgoal_tac "1 le n")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   514
apply (drule add_diff_inverse2, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   515
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   516
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   517
text{*Specialized for use in the proof below.*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   518
lemma multiset_funrestict:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   519
     "\<lbrakk>\<forall>a\<in>A. M ` a \<in> nat \<and> 0 < M ` a; Finite(A)\<rbrakk>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   520
      \<Longrightarrow> multiset(funrestrict(M, A - {a}))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   521
apply (simp add: multiset_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   522
apply (rule_tac x="A-{a}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   523
apply (auto intro: Finite_Diff funrestrict_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   524
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   525
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   526
lemma multiset_induct_aux:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   527
  assumes prem1: "!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   528
      and prem2: "!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   529
  shows
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   530
  "[| n \<in> nat; P(0) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   531
     ==> (\<forall>M. multiset(M)-->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   532
  (setsum(%x. $# mcount(M, x), {x \<in> mset_of(M). 0 < M`x}) = $# n) --> P(M))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   533
apply (erule nat_induct, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   534
apply (frule msize_eq_0_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   535
apply (auto simp add: mset_of_def multiset_def multiset_fun_iff msize_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   536
apply (subgoal_tac "setsum (%x. $# mcount (M, x), A) =$# succ (x) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   537
apply (drule setsum_succD, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   538
apply (case_tac "1 <M`a")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   539
apply (drule_tac [2] not_lt_imp_le)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   540
apply (simp_all add: nat_le_1_cases)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   541
apply (subgoal_tac "M= (M (a:=M`a #- 1)) (a:= (M (a:=M`a #- 1))`a #+ 1) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   542
apply (rule_tac [2] A = A and B = "%x. nat" and D = "%x. nat" in fun_extension)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   543
apply (rule_tac [3] update_type)+
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   544
apply (simp_all (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   545
 apply (rule_tac [2] impI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   546
 apply (rule_tac [2] succ_pred_eq_self [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   547
apply (simp_all (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   548
apply (rule subst, rule sym, blast, rule prem2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   549
apply (simp (no_asm) add: multiset_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   550
apply (rule_tac x = A in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   551
apply (force intro: update_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   552
apply (simp (no_asm_simp) add: mset_of_def mcount_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   553
apply (drule_tac x = "M (a := M ` a #- 1) " in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   554
apply (drule mp, drule_tac [2] mp, simp_all)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   555
apply (rule_tac x = A in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   556
apply (auto intro: update_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   557
apply (subgoal_tac "Finite ({x \<in> cons (a, A) . x\<noteq>a-->0<M`x}) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   558
prefer 2 apply (blast intro: Collect_subset [THEN subset_Finite] Finite_cons)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   559
apply (drule_tac A = "{x \<in> cons (a, A) . x\<noteq>a-->0<M`x}" in setsum_decr)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   560
apply (drule_tac x = M in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   561
apply (subgoal_tac "multiset (M) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   562
 prefer 2
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   563
 apply (simp add: multiset_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   564
 apply (rule_tac x = A in exI, force)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   565
apply (simp_all add: mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   566
apply (drule_tac psi = "\<forall>x \<in> A. ?u (x) " in asm_rl)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   567
apply (drule_tac x = a in bspec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   568
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   569
apply (subgoal_tac "cons (a, A) = A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   570
prefer 2 apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   571
apply simp
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   572
apply (subgoal_tac "M=cons (<a, M`a>, funrestrict (M, A-{a}))")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   573
 prefer 2
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   574
 apply (rule fun_cons_funrestrict_eq)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   575
 apply (subgoal_tac "cons (a, A-{a}) = A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   576
  apply force
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   577
  apply force
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   578
apply (rule_tac a = "cons (<a, 1>, funrestrict (M, A - {a}))" in ssubst)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   579
apply simp
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   580
apply (frule multiset_funrestict, assumption)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   581
apply (rule prem1, assumption)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   582
apply (simp add: mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   583
apply (drule_tac x = "funrestrict (M, A-{a}) " in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   584
apply (drule mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   585
apply (rule_tac x = "A-{a}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   586
apply (auto intro: Finite_Diff funrestrict_type simp add: funrestrict)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   587
apply (frule_tac A = A and M = M and a = a in setsum_decr3)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   588
apply (simp (no_asm_simp) add: multiset_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   589
apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   590
apply (simp (no_asm_simp) add: mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   591
apply (drule_tac b = "if ?u then ?v else ?w" in sym, simp_all)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   592
apply (subgoal_tac "{x \<in> A - {a} . 0 < funrestrict (M, A - {x}) ` x} = A - {a}")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   593
apply (auto intro!: setsum_cong simp add: zdiff_eq_iff zadd_commute multiset_def multiset_fun_iff mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   594
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   595
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   596
lemma multiset_induct2:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   597
  "[| multiset(M); P(0);
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   598
    (!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M)));
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   599
    (!!M b. [| multiset(M); b \<in> mset_of(M);  P(M) |] ==> P(M(b:= M`b #+ 1))) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   600
     ==> P(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   601
apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   602
apply (rule_tac [2] not_zneg_int_of)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   603
apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   604
apply (rule_tac [2] g_zpos_imp_setsum_zpos)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   605
prefer 2 apply (blast intro:  multiset_set_of_Finite Collect_subset [THEN subset_Finite])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   606
 prefer 2 apply (simp add: multiset_def multiset_fun_iff, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   607
apply (rule multiset_induct_aux [rule_format], auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   608
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   609
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   610
lemma munion_single_case1:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   611
     "[| multiset(M); a \<notin>mset_of(M) |] ==> M +# {#a#} = cons(<a, 1>, M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   612
apply (simp add: multiset_def msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   613
apply (auto simp add: munion_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   614
apply (unfold mset_of_def, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   615
apply (rule fun_extension, rule lam_type, simp_all)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   616
apply (auto simp add: multiset_fun_iff fun_extend_apply)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   617
apply (drule_tac c = a and b = 1 in fun_extend3)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   618
apply (auto simp add: cons_eq Un_commute [of _ "{a}"])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   619
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   620
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   621
lemma munion_single_case2:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   622
     "[| multiset(M); a \<in> mset_of(M) |] ==> M +# {#a#} = M(a:=M`a #+ 1)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   623
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   624
apply (auto simp add: munion_def multiset_fun_iff msingle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   625
apply (unfold mset_of_def, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   626
apply (subgoal_tac "A Un {a} = A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   627
apply (rule fun_extension)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   628
apply (auto dest: domain_type intro: lam_type update_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   629
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   630
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   631
(* Induction principle for multisets *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   632
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   633
lemma multiset_induct:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   634
  assumes M: "multiset(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   635
      and P0: "P(0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   636
      and step: "!!M a. [| multiset(M); P(M) |] ==> P(M +# {#a#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   637
  shows "P(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   638
apply (rule multiset_induct2 [OF M])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   639
apply (simp_all add: P0)
20898
113c9516a2d7 attribute symmetric: zero_var_indexes;
wenzelm
parents: 18415
diff changeset
   640
apply (frule_tac [2] a = b in munion_single_case2 [symmetric])
113c9516a2d7 attribute symmetric: zero_var_indexes;
wenzelm
parents: 18415
diff changeset
   641
apply (frule_tac a = a in munion_single_case1 [symmetric])
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   642
apply (auto intro: step)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   643
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   644
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   645
(** MCollect **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   646
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   647
lemma MCollect_multiset [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   648
     "multiset(M) ==> multiset({# x \<in> M. P(x)#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   649
apply (simp add: MCollect_def multiset_def mset_of_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   650
apply (rule_tac x = "{x \<in> A. P (x) }" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   651
apply (auto dest: CollectD1 [THEN [2] apply_type]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   652
            intro: Collect_subset [THEN subset_Finite] funrestrict_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   653
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   654
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   655
lemma mset_of_MCollect [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   656
     "multiset(M) ==> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   657
by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   658
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   659
lemma MCollect_mem_iff [iff]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   660
     "x \<in> mset_of({#x \<in> M. P(x)#}) <->  x \<in> mset_of(M) & P(x)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   661
by (simp add: MCollect_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   662
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   663
lemma mcount_MCollect [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   664
     "mcount({# x \<in> M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   665
by (simp add: mcount_def MCollect_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   666
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   667
lemma multiset_partition: "multiset(M) ==> M = {# x \<in> M. P(x) #} +# {# x \<in> M. ~ P(x) #}"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   668
by (simp add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   669
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   670
lemma natify_elem_is_self [simp]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   671
     "[| multiset(M); a \<in> mset_of(M) |] ==> natify(M`a) = M`a"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   672
by (auto simp add: multiset_def mset_of_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   673
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   674
(* and more algebraic laws on multisets *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   675
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   676
lemma munion_eq_conv_diff: "[| multiset(M); multiset(N) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   677
  ==>  (M +# {#a#} = N +# {#b#}) <->  (M = N & a = b |
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   678
       M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   679
apply (simp del: mcount_single add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   680
apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   681
apply (case_tac "a=b", auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   682
apply (drule_tac x = a in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   683
apply (drule_tac [2] x = b in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   684
apply (drule_tac [3] x = aa in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   685
apply (drule_tac [4] x = a in spec, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   686
apply (subgoal_tac [!] "mcount (N,a) :nat")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   687
apply (erule_tac [3] natE, erule natE, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   688
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   689
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   690
lemma melem_diff_single:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   691
"multiset(M) ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   692
  k \<in> mset_of(M -# {#a#}) <-> (k=a & 1 < mcount(M,a)) | (k\<noteq> a & k \<in> mset_of(M))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   693
apply (simp add: multiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   694
apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   695
apply (auto dest: domain_type intro: zero_less_diff [THEN iffD1]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   696
            simp add: multiset_fun_iff apply_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   697
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   698
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   699
apply (force intro!: lam_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   700
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   701
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   702
lemma munion_eq_conv_exist:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   703
"[| M \<in> Mult(A); N \<in> Mult(A) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   704
  ==> (M +# {#a#} = N +# {#b#}) <->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   705
      (M=N & a=b | (\<exists>K \<in> Mult(A). M= K +# {#b#} & N=K +# {#a#}))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   706
by (auto simp add: Mult_iff_multiset melem_diff_single munion_eq_conv_diff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   707
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   708
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   709
subsection{*Multiset Orderings*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   710
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   711
(* multiset on a domain A are finite functions from A to nat-{0} *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   712
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   713
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   714
(* multirel1 type *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   715
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   716
lemma multirel1_type: "multirel1(A, r) \<subseteq> Mult(A)*Mult(A)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   717
by (auto simp add: multirel1_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   718
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   719
lemma multirel1_0 [simp]: "multirel1(0, r) =0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   720
by (auto simp add: multirel1_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   721
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   722
lemma multirel1_iff:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   723
" <N, M> \<in> multirel1(A, r) <->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   724
  (\<exists>a. a \<in> A &
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   725
  (\<exists>M0. M0 \<in> Mult(A) & (\<exists>K. K \<in> Mult(A) &
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   726
   M=M0 +# {#a#} & N=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r))))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   727
by (auto simp add: multirel1_def Mult_iff_multiset Bex_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   728
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   729
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   730
text{*Monotonicity of @{term multirel1}*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   731
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   732
lemma multirel1_mono1: "A\<subseteq>B ==> multirel1(A, r)\<subseteq>multirel1(B, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   733
apply (auto simp add: multirel1_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   734
apply (auto simp add: Un_subset_iff Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   735
apply (rule_tac x = a in bexI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   736
apply (rule_tac x = M0 in bexI, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   737
apply (rule_tac x = K in bexI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   738
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   739
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   740
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   741
lemma multirel1_mono2: "r\<subseteq>s ==> multirel1(A,r)\<subseteq>multirel1(A, s)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   742
apply (simp add: multirel1_def, auto) 
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   743
apply (rule_tac x = a in bexI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   744
apply (rule_tac x = M0 in bexI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   745
apply (simp_all add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   746
apply (rule_tac x = K in bexI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   747
apply (simp_all add: Mult_iff_multiset, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   748
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   749
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   750
lemma multirel1_mono:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   751
     "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel1(A, r) \<subseteq> multirel1(B, s)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   752
apply (rule subset_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   753
apply (rule multirel1_mono1)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   754
apply (rule_tac [2] multirel1_mono2, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   755
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   756
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   757
subsection{* Toward the proof of well-foundedness of multirel1 *}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   758
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   759
lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   760
by (auto simp add: multirel1_def Mult_iff_multiset)
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
   761
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   762
lemma less_munion: "[| <N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A) |] ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   763
  (\<exists>M. <M, M0> \<in> multirel1(A, r) & N = M +# {#a#}) |
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   764
  (\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   765
apply (frule multirel1_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   766
apply (simp add: multirel1_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   767
apply (auto simp add: munion_eq_conv_exist)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   768
apply (rule_tac x="Ka +# K" in exI, auto, simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   769
apply (simp (no_asm_simp) add: munion_left_cancel munion_assoc)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   770
apply (auto simp add: munion_commute)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   771
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   772
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   773
lemma multirel1_base: "[| M \<in> Mult(A); a \<in> A |] ==> <M, M +# {#a#}> \<in> multirel1(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   774
apply (auto simp add: multirel1_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   775
apply (simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   776
apply (rule_tac x = a in exI, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   777
apply (rule_tac x = M in exI, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   778
apply (rule_tac x = 0 in exI, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   779
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   780
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   781
lemma acc_0: "acc(0)=0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   782
by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   783
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   784
lemma lemma1: "[| \<forall>b \<in> A. <b,a> \<in> r -->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   785
    (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r)));
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   786
    M0 \<in> acc(multirel1(A, r)); a \<in> A;
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   787
    \<forall>M. <M,M0> \<in> multirel1(A, r) --> M +# {#a#} \<in> acc(multirel1(A, r)) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   788
  ==> M0 +# {#a#} \<in> acc(multirel1(A, r))"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
   789
apply (subgoal_tac "M0 \<in> Mult(A) ")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   790
 prefer 2
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   791
 apply (erule acc.cases)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   792
 apply (erule fieldE)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   793
 apply (auto dest: multirel1_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   794
apply (rule accI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   795
apply (rename_tac "N")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   796
apply (drule less_munion, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   797
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   798
apply (erule_tac P = "\<forall>x \<in> mset_of (K) . <x, a> \<in> r" in rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   799
apply (erule_tac P = "mset_of (K) \<subseteq>A" in rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   800
apply (erule_tac M = K in multiset_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   801
(* three subgoals *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   802
(* subgoal 1: the induction base case *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   803
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   804
(* subgoal 2: the induction general case *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   805
apply (simp add: Ball_def Un_subset_iff, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   806
apply (drule_tac x = aa in spec, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   807
apply (subgoal_tac "aa \<in> A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   808
prefer 2 apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   809
apply (drule_tac x = "M0 +# M" and P =
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   810
       "%x. x \<in> acc(multirel1(A, r)) \<longrightarrow> ?Q(x)" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   811
apply (simp add: munion_assoc [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   812
(* subgoal 3: additional conditions *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   813
apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   814
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   815
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   816
lemma lemma2: "[| \<forall>b \<in> A. <b,a> \<in> r
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   817
   --> (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r)));
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   818
        M \<in> acc(multirel1(A, r)); a \<in> A|] ==> M +# {#a#} \<in> acc(multirel1(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   819
apply (erule acc_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   820
apply (blast intro: lemma1)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   821
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   822
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   823
lemma lemma3: "[| wf[A](r); a \<in> A |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   824
      ==> \<forall>M \<in> acc(multirel1(A, r)). M +# {#a#} \<in> acc(multirel1(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   825
apply (erule_tac a = a in wf_on_induct, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   826
apply (blast intro: lemma2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   827
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   828
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   829
lemma lemma4: "multiset(M) ==> mset_of(M)\<subseteq>A -->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   830
   wf[A](r) --> M \<in> field(multirel1(A, r)) --> M \<in> acc(multirel1(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   831
apply (erule multiset_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   832
(* proving the base case *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   833
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   834
apply (rule accI, force)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   835
apply (simp add: multirel1_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   836
(* Proving the general case *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   837
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   838
apply simp
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   839
apply (subgoal_tac "mset_of (M) \<subseteq>A")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   840
prefer 2 apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   841
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   842
apply (drule_tac a = a in lemma3, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   843
apply (subgoal_tac "M \<in> field (multirel1 (A,r))")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   844
apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   845
apply (rule multirel1_base [THEN fieldI1])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   846
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   847
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   848
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   849
lemma all_accessible: "[| wf[A](r); M \<in> Mult(A); A \<noteq> 0|] ==> M \<in> acc(multirel1(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   850
apply (erule not_emptyE)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   851
apply  (rule lemma4 [THEN mp, THEN mp, THEN mp])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   852
apply (rule_tac [4] multirel1_base [THEN fieldI1])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   853
apply  (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   854
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   855
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   856
lemma wf_on_multirel1: "wf[A](r) ==> wf[A-||>nat-{0}](multirel1(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   857
apply (case_tac "A=0")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   858
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   859
apply (rule wf_imp_wf_on)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   860
apply (rule wf_on_field_imp_wf)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   861
apply (simp (no_asm_simp) add: wf_on_0)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   862
apply (rule_tac A = "acc (multirel1 (A,r))" in wf_on_subset_A)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   863
apply (rule wf_on_acc)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   864
apply (blast intro: all_accessible)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   865
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   866
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   867
lemma wf_multirel1: "wf(r) ==>wf(multirel1(field(r), r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   868
apply (simp (no_asm_use) add: wf_iff_wf_on_field)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   869
apply (drule wf_on_multirel1)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   870
apply (rule_tac A = "field (r) -||> nat - {0}" in wf_on_subset_A)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   871
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   872
apply (rule field_rel_subset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   873
apply (rule multirel1_type)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   874
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   875
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   876
(** multirel **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   877
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   878
lemma multirel_type: "multirel(A, r) \<subseteq> Mult(A)*Mult(A)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   879
apply (simp add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   880
apply (rule trancl_type [THEN subset_trans])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   881
apply (auto dest: multirel1_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   882
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   883
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   884
(* Monotonicity of multirel *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   885
lemma multirel_mono:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   886
     "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel(A, r)\<subseteq>multirel(B,s)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   887
apply (simp add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   888
apply (rule trancl_mono)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   889
apply (rule multirel1_mono, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   890
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   891
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   892
(* Equivalence of multirel with the usual (closure-free) def *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   893
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   894
lemma add_diff_eq: "k \<in> nat ==> 0 < k --> n #+ k #- 1 = n #+ (k #- 1)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   895
by (erule nat_induct, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   896
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   897
lemma mdiff_union_single_conv: "[|a \<in> mset_of(J); multiset(I); multiset(J) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   898
   ==> I +# J -# {#a#} = I +# (J-# {#a#})"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   899
apply (simp (no_asm_simp) add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   900
apply (case_tac "a \<notin> mset_of (I) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   901
apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   902
apply (auto dest: domain_type simp add: add_diff_eq)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   903
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   904
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   905
lemma diff_add_commute: "[| n le m;  m \<in> nat; n \<in> nat; k \<in> nat |] ==> m #- n #+ k = m #+ k #- n"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   906
by (auto simp add: le_iff less_iff_succ_add)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   907
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   908
(* One direction *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   909
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   910
lemma multirel_implies_one_step:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   911
"<M,N> \<in> multirel(A, r) ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   912
     trans[A](r) -->
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   913
     (\<exists>I J K.
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   914
         I \<in> Mult(A) & J \<in> Mult(A) &  K \<in> Mult(A) &
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   915
         N = I +# J & M = I +# K & J \<noteq> 0 &
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   916
        (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   917
apply (simp add: multirel_def Ball_def Bex_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   918
apply (erule converse_trancl_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   919
apply (simp_all add: multirel1_iff Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   920
(* Two subgoals remain *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   921
(* Subgoal 1 *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   922
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   923
apply (rule_tac x = M0 in exI, force)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   924
(* Subgoal 2 *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   925
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   926
apply (case_tac "a \<in> mset_of (Ka) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   927
apply (rule_tac x = I in exI, simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   928
apply (rule_tac x = J in exI, simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   929
apply (rule_tac x = " (Ka -# {#a#}) +# K" in exI, simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   930
apply (simp_all add: Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   931
apply (simp (no_asm_simp) add: munion_assoc [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   932
apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   933
apply (simp add: mdiff_union_single_conv melem_diff_single, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   934
apply (erule disjE, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   935
apply (erule disjE, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   936
apply (drule_tac x = a and P = "%x. x :# Ka \<longrightarrow> ?Q(x)" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   937
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   938
apply (rule_tac x = xa in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   939
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   940
apply (blast dest: trans_onD)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   941
(* new we know that  a\<notin>mset_of(Ka) *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   942
apply (subgoal_tac "a :# I")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   943
apply (rule_tac x = "I-#{#a#}" in exI, simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   944
apply (rule_tac x = "J+#{#a#}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   945
apply (simp (no_asm_simp) add: Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   946
apply (rule_tac x = "Ka +# K" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   947
apply (simp (no_asm_simp) add: Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   948
apply (rule conjI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   949
apply (simp (no_asm_simp) add: multiset_equality mcount_elem [THEN succ_pred_eq_self])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   950
apply (rule conjI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   951
apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   952
apply (simp add: mdiff_union_inverse2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   953
apply (simp_all (no_asm_simp) add: multiset_equality)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   954
apply (rule diff_add_commute [symmetric])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   955
apply (auto intro: mcount_elem)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   956
apply (subgoal_tac "a \<in> mset_of (I +# Ka) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   957
apply (drule_tac [2] sym, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   958
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   959
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   960
lemma melem_imp_eq_diff_union [simp]: "[| a \<in> mset_of(M); multiset(M) |] ==> M -# {#a#} +# {#a#} = M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   961
by (simp add: multiset_equality mcount_elem [THEN succ_pred_eq_self])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   962
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   963
lemma msize_eq_succ_imp_eq_union:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   964
     "[| msize(M)=$# succ(n); M \<in> Mult(A); n \<in> nat |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   965
      ==> \<exists>a N. M = N +# {#a#} & N \<in> Mult(A) & a \<in> A"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   966
apply (drule msize_eq_succ_imp_elem, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   967
apply (rule_tac x = a in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   968
apply (rule_tac x = "M -# {#a#}" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   969
apply (frule Mult_into_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   970
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   971
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   972
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   973
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   974
(* The second direction *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   975
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   976
lemma one_step_implies_multirel_lemma [rule_format (no_asm)]:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   977
"n \<in> nat ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   978
   (\<forall>I J K.
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   979
    I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   980
   (msize(J) = $# n & J \<noteq>0 &  (\<forall>k \<in> mset_of(K).  \<exists>j \<in> mset_of(J). <k, j> \<in> r))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   981
    --> <I +# K, I +# J> \<in> multirel(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   982
apply (simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   983
apply (erule nat_induct, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   984
apply (drule_tac M = J in msize_eq_0_iff, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   985
(* one subgoal remains *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   986
apply (subgoal_tac "msize (J) =$# succ (x) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   987
 prefer 2 apply simp
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   988
apply (frule_tac A = A in msize_eq_succ_imp_eq_union)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   989
apply (simp_all add: Mult_iff_multiset, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   990
apply (rename_tac "J'", simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   991
apply (case_tac "J' = 0")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   992
apply (simp add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   993
apply (rule r_into_trancl, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   994
apply (simp add: multirel1_iff Mult_iff_multiset, force)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   995
(*Now we know J' \<noteq>  0*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   996
apply (drule sym, rotate_tac -1, simp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   997
apply (erule_tac V = "$# x = msize (J') " in thin_rl)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   998
apply (frule_tac M = K and P = "%x. <x,a> \<in> r" in multiset_partition)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
   999
apply (erule_tac P = "\<forall>k \<in> mset_of (K) . ?P (k) " in rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1000
apply (erule ssubst)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1001
apply (simp add: Ball_def, auto)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1002
apply (subgoal_tac "< (I +# {# x \<in> K. <x, a> \<in> r#}) +# {# x \<in> K. <x, a> \<notin> r#}, (I +# {# x \<in> K. <x, a> \<in> r#}) +# J'> \<in> multirel(A, r) ")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1003
 prefer 2
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1004
 apply (drule_tac x = "I +# {# x \<in> K. <x, a> \<in> r#}" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1005
 apply (rotate_tac -1)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1006
 apply (drule_tac x = "J'" in spec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1007
 apply (rotate_tac -1)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1008
 apply (drule_tac x = "{# x \<in> K. <x, a> \<notin> r#}" in spec, simp) apply blast
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1009
apply (simp add: munion_assoc [symmetric] multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1010
apply (rule_tac b = "I +# {# x \<in> K. <x, a> \<in> r#} +# J'" in trancl_trans, blast)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1011
apply (rule r_into_trancl)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1012
apply (simp add: multirel1_iff Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1013
apply (rule_tac x = a in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1014
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1015
apply (rule_tac x = "I +# J'" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1016
apply (auto simp add: munion_ac Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1017
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1018
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1019
lemma one_step_implies_multirel:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1020
     "[| J \<noteq> 0;  \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r;
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1021
         I \<in> Mult(A); J \<in> Mult(A); K \<in> Mult(A) |]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1022
      ==> <I+#K, I+#J> \<in> multirel(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1023
apply (subgoal_tac "multiset (J) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1024
 prefer 2 apply (simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1025
apply (frule_tac M = J in msize_int_of_nat)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1026
apply (auto intro: one_step_implies_multirel_lemma)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1027
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1028
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1029
(** Proving that multisets are partially ordered **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1030
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1031
(*irreflexivity*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1032
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1033
lemma multirel_irrefl_lemma:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1034
     "Finite(A) ==> part_ord(A, r) --> (\<forall>x \<in> A. \<exists>y \<in> A. <x,y> \<in> r) -->A=0"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1035
apply (erule Finite_induct)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1036
apply (auto dest: subset_consI [THEN [2] part_ord_subset])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1037
apply (auto simp add: part_ord_def irrefl_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1038
apply (drule_tac x = xa in bspec)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1039
apply (drule_tac [2] a = xa and b = x in trans_onD, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1040
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1041
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1042
lemma irrefl_on_multirel:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1043
     "part_ord(A, r) ==> irrefl(Mult(A), multirel(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1044
apply (simp add: irrefl_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1045
apply (subgoal_tac "trans[A](r) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1046
 prefer 2 apply (simp add: part_ord_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1047
apply (drule multirel_implies_one_step, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1048
apply (simp add: Mult_iff_multiset, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1049
apply (subgoal_tac "Finite (mset_of (K))")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1050
apply (frule_tac r = r in multirel_irrefl_lemma)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1051
apply (frule_tac B = "mset_of (K) " in part_ord_subset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1052
apply simp_all
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1053
apply (auto simp add: multiset_def mset_of_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1054
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1055
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1056
lemma trans_on_multirel: "trans[Mult(A)](multirel(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1057
apply (simp add: multirel_def trans_on_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1058
apply (blast intro: trancl_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1059
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1060
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1061
lemma multirel_trans:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1062
 "[| <M, N> \<in> multirel(A, r); <N, K> \<in> multirel(A, r) |] ==>  <M, K> \<in> multirel(A,r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1063
apply (simp add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1064
apply (blast intro: trancl_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1065
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1066
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1067
lemma trans_multirel: "trans(multirel(A,r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1068
apply (simp add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1069
apply (rule trans_trancl)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1070
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1071
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1072
lemma part_ord_multirel: "part_ord(A,r) ==> part_ord(Mult(A), multirel(A, r))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1073
apply (simp (no_asm) add: part_ord_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1074
apply (blast intro: irrefl_on_multirel trans_on_multirel)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1075
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1076
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1077
(** Monotonicity of multiset union **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1078
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1079
lemma munion_multirel1_mono:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1080
"[|<M,N> \<in> multirel1(A, r); K \<in> Mult(A) |] ==> <K +# M, K +# N> \<in> multirel1(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1081
apply (frule multirel1_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1082
apply (auto simp add: multirel1_iff Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1083
apply (rule_tac x = a in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1084
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1085
apply (rule_tac x = "K+#M0" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1086
apply (simp (no_asm_simp) add: Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1087
apply (rule_tac x = Ka in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1088
apply (simp (no_asm_simp) add: munion_assoc)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1089
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1090
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1091
lemma munion_multirel_mono2:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1092
 "[| <M, N> \<in> multirel(A, r); K \<in> Mult(A) |]==><K +# M, K +# N> \<in> multirel(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1093
apply (frule multirel_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1094
apply (simp (no_asm_use) add: multirel_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1095
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1096
apply (drule_tac psi = "<M,N> \<in> multirel1 (A, r) ^+" in asm_rl)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1097
apply (erule rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1098
apply (erule rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1099
apply (erule rev_mp)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1100
apply (erule trancl_induct, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1101
apply (blast intro: munion_multirel1_mono r_into_trancl, clarify)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1102
apply (subgoal_tac "y \<in> Mult(A) ")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1103
 prefer 2
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1104
 apply (blast dest: multirel_type [unfolded multirel_def, THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1105
apply (subgoal_tac "<K +# y, K +# z> \<in> multirel1 (A, r) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1106
prefer 2 apply (blast intro: munion_multirel1_mono)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1107
apply (blast intro: r_into_trancl trancl_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1108
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1109
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1110
lemma munion_multirel_mono1:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1111
     "[|<M, N> \<in> multirel(A, r); K \<in> Mult(A)|] ==> <M +# K, N +# K> \<in> multirel(A, r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1112
apply (frule multirel_type [THEN subsetD])
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1113
apply (rule_tac P = "%x. <x,?u> \<in> multirel(A, r) " in munion_commute [THEN subst])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1114
apply (subst munion_commute [of N])
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1115
apply (rule munion_multirel_mono2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1116
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1117
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1118
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1119
lemma munion_multirel_mono:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1120
     "[|<M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)|]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1121
      ==> <M +# N, K +# L> \<in> multirel(A, r)"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1122
apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1123
prefer 2 apply (blast dest: multirel_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1124
apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1125
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1126
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1127
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1128
subsection{*Ordinal Multisets*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1129
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1130
(* A \<subseteq> B ==>  field(Memrel(A)) \<subseteq> field(Memrel(B)) *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1131
lemmas field_Memrel_mono = Memrel_mono [THEN field_mono, standard]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1132
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1133
(*
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1134
[| Aa \<subseteq> Ba; A \<subseteq> B |] ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1135
multirel(field(Memrel(Aa)), Memrel(A))\<subseteq> multirel(field(Memrel(Ba)), Memrel(B))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1136
*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1137
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1138
lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1139
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1140
lemma omultiset_is_multiset [simp]: "omultiset(M) ==> multiset(M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1141
apply (simp add: omultiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1142
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1143
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1144
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1145
lemma munion_omultiset [simp]: "[| omultiset(M); omultiset(N) |] ==> omultiset(M +# N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1146
apply (simp add: omultiset_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1147
apply (rule_tac x = "i Un ia" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1148
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1149
apply (blast intro: field_Memrel_mono)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1150
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1151
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1152
lemma mdiff_omultiset [simp]: "omultiset(M) ==> omultiset(M -# N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1153
apply (simp add: omultiset_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1154
apply (simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1155
apply (rule_tac x = i in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1156
apply (simp (no_asm_simp))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1157
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1158
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1159
(** Proving that Memrel is a partial order **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1160
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1161
lemma irrefl_Memrel: "Ord(i) ==> irrefl(field(Memrel(i)), Memrel(i))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1162
apply (rule irreflI, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1163
apply (subgoal_tac "Ord (x) ")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1164
prefer 2 apply (blast intro: Ord_in_Ord)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1165
apply (drule_tac i = x in ltI [THEN lt_irrefl], auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1166
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1167
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1168
lemma trans_iff_trans_on: "trans(r) <-> trans[field(r)](r)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1169
by (simp add: trans_on_def trans_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1170
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1171
lemma part_ord_Memrel: "Ord(i) ==>part_ord(field(Memrel(i)), Memrel(i))"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1172
apply (simp add: part_ord_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1173
apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1174
apply (blast intro: trans_Memrel irrefl_Memrel)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1175
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1176
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1177
(*
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1178
  Ord(i) ==>
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1179
  part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i)))
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1180
*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1181
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1182
lemmas part_ord_mless = part_ord_Memrel [THEN part_ord_multirel, standard]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1183
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1184
(*irreflexivity*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1185
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1186
lemma mless_not_refl: "~(M <# M)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1187
apply (simp add: mless_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1188
apply (frule multirel_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1189
apply (drule part_ord_mless)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1190
apply (simp add: part_ord_def irrefl_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1191
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1192
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1193
(* N<N ==> R *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1194
lemmas mless_irrefl = mless_not_refl [THEN notE, standard, elim!]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1195
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1196
(*transitivity*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1197
lemma mless_trans: "[| K <# M; M <# N |] ==> K <# N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1198
apply (simp add: mless_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1199
apply (rule_tac x = "i Un ia" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1200
apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1201
                   multirel_Memrel_mono [OF Un_upper2 Un_upper2, THEN subsetD]
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1202
        intro: multirel_trans Ord_Un)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1203
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1204
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1205
(*asymmetry*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1206
lemma mless_not_sym: "M <# N ==> ~ N <# M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1207
apply clarify
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1208
apply (rule mless_not_refl [THEN notE])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1209
apply (erule mless_trans, assumption)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1210
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1211
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1212
lemma mless_asym: "[| M <# N; ~P ==> N <# M |] ==> P"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1213
by (blast dest: mless_not_sym)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1214
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1215
lemma mle_refl [simp]: "omultiset(M) ==> M <#= M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1216
by (simp add: mle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1217
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1218
(*anti-symmetry*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1219
lemma mle_antisym:
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1220
     "[| M <#= N;  N <#= M |] ==> M = N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1221
apply (simp add: mle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1222
apply (blast dest: mless_not_sym)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1223
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1224
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1225
(*transitivity*)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1226
lemma mle_trans: "[| K <#= M; M <#= N |] ==> K <#= N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1227
apply (simp add: mle_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1228
apply (blast intro: mless_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1229
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1230
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1231
lemma mless_le_iff: "M <# N <-> (M <#= N & M \<noteq> N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1232
by (simp add: mle_def, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1233
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1234
(** Monotonicity of mless **)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1235
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1236
lemma munion_less_mono2: "[| M <# N; omultiset(K) |] ==> K +# M <# K +# N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1237
apply (simp add: mless_def omultiset_def, clarify)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1238
apply (rule_tac x = "i Un ia" in exI)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1239
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1240
apply (rule munion_multirel_mono2)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1241
 apply (blast intro: multirel_Memrel_mono [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1242
apply (simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1243
apply (blast intro: field_Memrel_mono [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1244
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1245
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1246
lemma munion_less_mono1: "[| M <# N; omultiset(K) |] ==> M +# K <# N +# K"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1247
by (force dest: munion_less_mono2 simp add: munion_commute)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1248
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1249
lemma mless_imp_omultiset: "M <# N ==> omultiset(M) & omultiset(N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1250
by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD])
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1251
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1252
lemma munion_less_mono: "[| M <# K; N <# L |] ==> M +# N <# K +# L"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1253
apply (frule_tac M = M in mless_imp_omultiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1254
apply (frule_tac M = N in mless_imp_omultiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1255
apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1256
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1257
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1258
(* <#= *)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1259
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1260
lemma mle_imp_omultiset: "M <#= N ==> omultiset(M) & omultiset(N)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1261
by (auto simp add: mle_def mless_imp_omultiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1262
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1263
lemma mle_mono: "[| M <#= K;  N <#= L |] ==> M +# N <#= K +# L"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1264
apply (frule_tac M = M in mle_imp_omultiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1265
apply (frule_tac M = N in mle_imp_omultiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1266
apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1267
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1268
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1269
lemma omultiset_0 [iff]: "omultiset(0)"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1270
by (auto simp add: omultiset_def Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1271
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1272
lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1273
apply (simp add: mle_def mless_def)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1274
apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1275
 prefer 2 apply (simp add: omultiset_def)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1276
apply (case_tac "M=0", simp_all, clarify)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15201
diff changeset
  1277
apply (subgoal_tac "<0 +# 0, 0 +# M> \<in> multirel(field (Memrel(i)), Memrel(i))")
15201
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1278
apply (rule_tac [2] one_step_implies_multirel)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1279
apply (auto simp add: Mult_iff_multiset)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1280
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1281
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1282
lemma munion_upper1: "[| omultiset(M); omultiset(N) |] ==> M <#= M +# N"
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1283
apply (subgoal_tac "M +# 0 <#= M +# N")
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1284
apply (rule_tac [2] mle_mono, auto)
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1285
done
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1286
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1287
ML
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1288
{*
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1289
val munion_ac = thms "munion_ac";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1290
val funrestrict_subset = thm "funrestrict_subset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1291
val funrestrict_type = thm "funrestrict_type";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1292
val funrestrict_type2 = thm "funrestrict_type2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1293
val funrestrict = thm "funrestrict";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1294
val funrestrict_empty = thm "funrestrict_empty";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1295
val domain_funrestrict = thm "domain_funrestrict";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1296
val fun_cons_funrestrict_eq = thm "fun_cons_funrestrict_eq";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1297
val multiset_fun_iff = thm "multiset_fun_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1298
val multiset_into_Mult = thm "multiset_into_Mult";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1299
val Mult_into_multiset = thm "Mult_into_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1300
val Mult_iff_multiset = thm "Mult_iff_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1301
val multiset_iff_Mult_mset_of = thm "multiset_iff_Mult_mset_of";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1302
val multiset_0 = thm "multiset_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1303
val multiset_set_of_Finite = thm "multiset_set_of_Finite";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1304
val mset_of_0 = thm "mset_of_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1305
val mset_is_0_iff = thm "mset_is_0_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1306
val mset_of_single = thm "mset_of_single";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1307
val mset_of_union = thm "mset_of_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1308
val mset_of_diff = thm "mset_of_diff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1309
val msingle_not_0 = thm "msingle_not_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1310
val msingle_eq_iff = thm "msingle_eq_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1311
val msingle_multiset = thm "msingle_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1312
val Collect_Finite = thms "Collect_Finite";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1313
val normalize_idem = thm "normalize_idem";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1314
val normalize_multiset = thm "normalize_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1315
val multiset_normalize = thm "multiset_normalize";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1316
val munion_multiset = thm "munion_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1317
val mdiff_multiset = thm "mdiff_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1318
val munion_0 = thm "munion_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1319
val munion_commute = thm "munion_commute";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1320
val munion_assoc = thm "munion_assoc";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1321
val munion_lcommute = thm "munion_lcommute";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1322
val mdiff_self_eq_0 = thm "mdiff_self_eq_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1323
val mdiff_0 = thm "mdiff_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1324
val mdiff_0_right = thm "mdiff_0_right";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1325
val mdiff_union_inverse2 = thm "mdiff_union_inverse2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1326
val mcount_type = thm "mcount_type";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1327
val mcount_0 = thm "mcount_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1328
val mcount_single = thm "mcount_single";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1329
val mcount_union = thm "mcount_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1330
val mcount_diff = thm "mcount_diff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1331
val mcount_elem = thm "mcount_elem";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1332
val msize_0 = thm "msize_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1333
val msize_single = thm "msize_single";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1334
val msize_type = thm "msize_type";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1335
val msize_zpositive = thm "msize_zpositive";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1336
val msize_int_of_nat = thm "msize_int_of_nat";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1337
val not_empty_multiset_imp_exist = thm "not_empty_multiset_imp_exist";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1338
val msize_eq_0_iff = thm "msize_eq_0_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1339
val setsum_mcount_Int = thm "setsum_mcount_Int";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1340
val msize_union = thm "msize_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1341
val msize_eq_succ_imp_elem = thm "msize_eq_succ_imp_elem";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1342
val multiset_equality = thm "multiset_equality";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1343
val munion_eq_0_iff = thm "munion_eq_0_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1344
val empty_eq_munion_iff = thm "empty_eq_munion_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1345
val munion_right_cancel = thm "munion_right_cancel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1346
val munion_left_cancel = thm "munion_left_cancel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1347
val nat_add_eq_1_cases = thm "nat_add_eq_1_cases";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1348
val munion_is_single = thm "munion_is_single";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1349
val msingle_is_union = thm "msingle_is_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1350
val setsum_decr = thm "setsum_decr";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1351
val setsum_decr2 = thm "setsum_decr2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1352
val setsum_decr3 = thm "setsum_decr3";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1353
val nat_le_1_cases = thm "nat_le_1_cases";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1354
val succ_pred_eq_self = thm "succ_pred_eq_self";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1355
val multiset_funrestict = thm "multiset_funrestict";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1356
val multiset_induct_aux = thm "multiset_induct_aux";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1357
val multiset_induct2 = thm "multiset_induct2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1358
val munion_single_case1 = thm "munion_single_case1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1359
val munion_single_case2 = thm "munion_single_case2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1360
val multiset_induct = thm "multiset_induct";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1361
val MCollect_multiset = thm "MCollect_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1362
val mset_of_MCollect = thm "mset_of_MCollect";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1363
val MCollect_mem_iff = thm "MCollect_mem_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1364
val mcount_MCollect = thm "mcount_MCollect";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1365
val multiset_partition = thm "multiset_partition";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1366
val natify_elem_is_self = thm "natify_elem_is_self";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1367
val munion_eq_conv_diff = thm "munion_eq_conv_diff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1368
val melem_diff_single = thm "melem_diff_single";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1369
val munion_eq_conv_exist = thm "munion_eq_conv_exist";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1370
val multirel1_type = thm "multirel1_type";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1371
val multirel1_0 = thm "multirel1_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1372
val multirel1_iff = thm "multirel1_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1373
val multirel1_mono1 = thm "multirel1_mono1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1374
val multirel1_mono2 = thm "multirel1_mono2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1375
val multirel1_mono = thm "multirel1_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1376
val not_less_0 = thm "not_less_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1377
val less_munion = thm "less_munion";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1378
val multirel1_base = thm "multirel1_base";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1379
val acc_0 = thm "acc_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1380
val all_accessible = thm "all_accessible";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1381
val wf_on_multirel1 = thm "wf_on_multirel1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1382
val wf_multirel1 = thm "wf_multirel1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1383
val multirel_type = thm "multirel_type";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1384
val multirel_mono = thm "multirel_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1385
val add_diff_eq = thm "add_diff_eq";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1386
val mdiff_union_single_conv = thm "mdiff_union_single_conv";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1387
val diff_add_commute = thm "diff_add_commute";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1388
val multirel_implies_one_step = thm "multirel_implies_one_step";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1389
val melem_imp_eq_diff_union = thm "melem_imp_eq_diff_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1390
val msize_eq_succ_imp_eq_union = thm "msize_eq_succ_imp_eq_union";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1391
val one_step_implies_multirel = thm "one_step_implies_multirel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1392
val irrefl_on_multirel = thm "irrefl_on_multirel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1393
val trans_on_multirel = thm "trans_on_multirel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1394
val multirel_trans = thm "multirel_trans";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1395
val trans_multirel = thm "trans_multirel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1396
val part_ord_multirel = thm "part_ord_multirel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1397
val munion_multirel1_mono = thm "munion_multirel1_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1398
val munion_multirel_mono2 = thm "munion_multirel_mono2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1399
val munion_multirel_mono1 = thm "munion_multirel_mono1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1400
val munion_multirel_mono = thm "munion_multirel_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1401
val field_Memrel_mono = thms "field_Memrel_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1402
val multirel_Memrel_mono = thms "multirel_Memrel_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1403
val omultiset_is_multiset = thm "omultiset_is_multiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1404
val munion_omultiset = thm "munion_omultiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1405
val mdiff_omultiset = thm "mdiff_omultiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1406
val irrefl_Memrel = thm "irrefl_Memrel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1407
val trans_iff_trans_on = thm "trans_iff_trans_on";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1408
val part_ord_Memrel = thm "part_ord_Memrel";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1409
val part_ord_mless = thms "part_ord_mless";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1410
val mless_not_refl = thm "mless_not_refl";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1411
val mless_irrefl = thms "mless_irrefl";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1412
val mless_trans = thm "mless_trans";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1413
val mless_not_sym = thm "mless_not_sym";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1414
val mless_asym = thm "mless_asym";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1415
val mle_refl = thm "mle_refl";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1416
val mle_antisym = thm "mle_antisym";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1417
val mle_trans = thm "mle_trans";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1418
val mless_le_iff = thm "mless_le_iff";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1419
val munion_less_mono2 = thm "munion_less_mono2";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1420
val munion_less_mono1 = thm "munion_less_mono1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1421
val mless_imp_omultiset = thm "mless_imp_omultiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1422
val munion_less_mono = thm "munion_less_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1423
val mle_imp_omultiset = thm "mle_imp_omultiset";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1424
val mle_mono = thm "mle_mono";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1425
val omultiset_0 = thm "omultiset_0";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1426
val empty_leI = thm "empty_leI";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1427
val munion_upper1 = thm "munion_upper1";
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1428
*}
d73f9d49d835 converted ZF/Induct/Multiset to Isar script
paulson
parents: 14046
diff changeset
  1429
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
  1430
end