src/HOL/IOA/NTP/Multiset.ML
changeset 1051 4fcd0638e61d
child 1266 3ae9fe3c0f68
equal deleted inserted replaced
1050:0c36c6a52a1d 1051:4fcd0638e61d
       
     1 (*  Title:      HOL/IOA/NTP/Multiset.ML
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow & Konrad Slind
       
     4     Copyright   1994  TU Muenchen
       
     5 
       
     6 Axiomatic multisets.
       
     7 Should be done as a subtype and moved to a global place.
       
     8 *)
       
     9 
       
    10 goalw Multiset.thy [Multiset.count_def, Multiset.countm_empty_def]
       
    11    "count {|} x = 0";
       
    12  by (rtac refl 1);
       
    13 qed "count_empty";
       
    14 
       
    15 goal Multiset.thy 
       
    16      "count (addm M x) y = (if y=x then Suc(count M y) else count M y)";
       
    17   by (asm_simp_tac (arith_ss addsimps 
       
    18                     [Multiset.count_def,Multiset.countm_nonempty_def]
       
    19                     setloop (split_tac [expand_if])) 1);
       
    20 qed "count_addm_simp";
       
    21 
       
    22 goal Multiset.thy "count M y <= count (addm M x) y";
       
    23   by (simp_tac (arith_ss addsimps [count_addm_simp]
       
    24                          setloop (split_tac [expand_if])) 1);
       
    25   by (rtac impI 1);
       
    26   by (rtac (le_refl RS (leq_suc RS mp)) 1);
       
    27 qed "count_leq_addm";
       
    28 
       
    29 goalw Multiset.thy [Multiset.count_def] 
       
    30      "count (delm M x) y = (if y=x then pred(count M y) else count M y)";
       
    31   by (res_inst_tac [("M","M")] Multiset.induction 1);
       
    32   by (asm_simp_tac (arith_ss 
       
    33                    addsimps [Multiset.delm_empty_def,Multiset.countm_empty_def]
       
    34                    setloop (split_tac [expand_if])) 1);
       
    35   by (asm_full_simp_tac (arith_ss 
       
    36                          addsimps [Multiset.delm_nonempty_def,
       
    37                                    Multiset.countm_nonempty_def]
       
    38                          setloop (split_tac [expand_if])) 1);
       
    39   by (safe_tac HOL_cs);
       
    40   by (asm_full_simp_tac HOL_ss 1);
       
    41 qed "count_delm_simp";
       
    42 
       
    43 goal Multiset.thy "!!M. (!x. P(x) --> Q(x)) ==> (countm M P <= countm M Q)";
       
    44   by (res_inst_tac [("M","M")] Multiset.induction 1);
       
    45   by (simp_tac (arith_ss addsimps [Multiset.countm_empty_def]) 1);
       
    46   by (simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def]) 1);
       
    47   by (etac (less_eq_add_cong RS mp RS mp) 1);
       
    48   by (asm_full_simp_tac (arith_ss addsimps [le_eq_less_or_eq]
       
    49                                   setloop (split_tac [expand_if])) 1);
       
    50 qed "countm_props";
       
    51 
       
    52 goal Multiset.thy "!!P. ~P(obj) ==> countm M P = countm (delm M obj) P";
       
    53   by (res_inst_tac [("M","M")] Multiset.induction 1);
       
    54   by (simp_tac (arith_ss addsimps [Multiset.delm_empty_def,
       
    55                                    Multiset.countm_empty_def]) 1);
       
    56   by (asm_simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def,
       
    57                                       Multiset.delm_nonempty_def]
       
    58                              setloop (split_tac [expand_if])) 1);
       
    59 qed "countm_spurious_delm";
       
    60 
       
    61 
       
    62 goal Multiset.thy "!!P. P(x) ==> 0<count M x --> 0<countm M P";
       
    63   by (res_inst_tac [("M","M")] Multiset.induction 1);
       
    64   by (simp_tac (arith_ss addsimps 
       
    65                           [Multiset.delm_empty_def,Multiset.count_def,
       
    66                            Multiset.countm_empty_def]) 1);
       
    67   by (asm_simp_tac (arith_ss addsimps 
       
    68                        [Multiset.count_def,Multiset.delm_nonempty_def,
       
    69                         Multiset.countm_nonempty_def]
       
    70                     setloop (split_tac [expand_if])) 1);
       
    71 val pos_count_imp_pos_countm = store_thm("pos_count_imp_pos_countm", standard(result() RS mp));
       
    72 
       
    73 goal Multiset.thy
       
    74    "!!P. P(x) ==> 0<count M x --> countm (delm M x) P = pred (countm M P)";
       
    75   by (res_inst_tac [("M","M")] Multiset.induction 1);
       
    76   by (simp_tac (arith_ss addsimps 
       
    77                           [Multiset.delm_empty_def,
       
    78                            Multiset.countm_empty_def]) 1);
       
    79   by (asm_simp_tac (arith_ss addsimps 
       
    80                       [eq_sym_conv,count_addm_simp,Multiset.delm_nonempty_def,
       
    81                        Multiset.countm_nonempty_def,pos_count_imp_pos_countm,
       
    82                        suc_pred_id]
       
    83                     setloop (split_tac [expand_if])) 1);
       
    84 qed "countm_done_delm";