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