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