IOA/example/Multiset.ML
changeset 171 16c4ea954511
parent 168 44ff2275d44f
equal deleted inserted replaced
170:3a8d722fd3ff 171:16c4ea954511
     8 *)
     8 *)
     9 
     9 
    10 goalw Multiset.thy [Multiset.count_def, Multiset.countm_empty_def]
    10 goalw Multiset.thy [Multiset.count_def, Multiset.countm_empty_def]
    11    "count({|},x) = 0";
    11    "count({|},x) = 0";
    12  by (rtac refl 1);
    12  by (rtac refl 1);
    13 val count_empty = result();
    13 qed "count_empty";
    14 
    14 
    15 goal Multiset.thy 
    15 goal Multiset.thy 
    16      "count(addm(M,x),y) = if(y=x,Suc(count(M,y)),count(M,y))";
    16      "count(addm(M,x),y) = if(y=x,Suc(count(M,y)),count(M,y))";
    17   by (asm_simp_tac (arith_ss addsimps 
    17   by (asm_simp_tac (arith_ss addsimps 
    18                     [Multiset.count_def,Multiset.countm_nonempty_def]
    18                     [Multiset.count_def,Multiset.countm_nonempty_def]
    19                     setloop (split_tac [expand_if])) 1);
    19                     setloop (split_tac [expand_if])) 1);
    20 val count_addm_simp = result();
    20 qed "count_addm_simp";
    21 
    21 
    22 goal Multiset.thy "count(M,y) <= count(addm(M,x),y)";
    22 goal Multiset.thy "count(M,y) <= count(addm(M,x),y)";
    23   by (simp_tac (arith_ss addsimps [count_addm_simp]
    23   by (simp_tac (arith_ss addsimps [count_addm_simp]
    24                          setloop (split_tac [expand_if])) 1);
    24                          setloop (split_tac [expand_if])) 1);
    25   by (rtac impI 1);
    25   by (rtac impI 1);
    26   by (rtac (le_refl RS (leq_suc RS mp)) 1);
    26   by (rtac (le_refl RS (leq_suc RS mp)) 1);
    27 val count_leq_addm = result();
    27 qed "count_leq_addm";
    28 
    28 
    29 goalw Multiset.thy [Multiset.count_def] 
    29 goalw Multiset.thy [Multiset.count_def] 
    30      "count(delm(M,x),y) = if(y=x,pred(count(M,y)),count(M,y))";
    30      "count(delm(M,x),y) = if(y=x,pred(count(M,y)),count(M,y))";
    31   by (res_inst_tac [("M","M")] Multiset.induction 1);
    31   by (res_inst_tac [("M","M")] Multiset.induction 1);
    32   by (asm_simp_tac (arith_ss 
    32   by (asm_simp_tac (arith_ss 
    36                          addsimps [Multiset.delm_nonempty_def,
    36                          addsimps [Multiset.delm_nonempty_def,
    37                                    Multiset.countm_nonempty_def]
    37                                    Multiset.countm_nonempty_def]
    38                          setloop (split_tac [expand_if])) 1);
    38                          setloop (split_tac [expand_if])) 1);
    39   by (safe_tac HOL_cs);
    39   by (safe_tac HOL_cs);
    40   by (asm_full_simp_tac HOL_ss 1);
    40   by (asm_full_simp_tac HOL_ss 1);
    41 val count_delm_simp = result();
    41 qed "count_delm_simp";
    42 
    42 
    43 goal Multiset.thy "!!M. (!x. P(x) --> Q(x)) ==> (countm(M,P) <= countm(M,Q))";
    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);
    44   by (res_inst_tac [("M","M")] Multiset.induction 1);
    45   by (simp_tac (arith_ss addsimps [Multiset.countm_empty_def]) 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);
    46   by (simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def]) 1);
    47   by (etac (less_eq_add_cong RS mp RS mp) 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]
    48   by (asm_full_simp_tac (arith_ss addsimps [le_eq_less_or_eq]
    49                                   setloop (split_tac [expand_if])) 1);
    49                                   setloop (split_tac [expand_if])) 1);
    50 val countm_props = result();
    50 qed "countm_props";
    51 
    51 
    52 goal Multiset.thy "!!P. ~P(obj) ==> countm(M,P) = countm(delm(M,obj),P)";
    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);
    53   by (res_inst_tac [("M","M")] Multiset.induction 1);
    54   by (simp_tac (arith_ss addsimps [Multiset.delm_empty_def,
    54   by (simp_tac (arith_ss addsimps [Multiset.delm_empty_def,
    55                                    Multiset.countm_empty_def]) 1);
    55                                    Multiset.countm_empty_def]) 1);
    56   by (asm_simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def,
    56   by (asm_simp_tac (arith_ss addsimps[Multiset.countm_nonempty_def,
    57                                       Multiset.delm_nonempty_def]
    57                                       Multiset.delm_nonempty_def]
    58                              setloop (split_tac [expand_if])) 1);
    58                              setloop (split_tac [expand_if])) 1);
    59 val countm_spurious_delm = result();
    59 qed "countm_spurious_delm";
    60 
    60 
    61 
    61 
    62 goal Multiset.thy "!!P. P(x) ==> 0<count(M,x) --> 0<countm(M,P)";
    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);
    63   by (res_inst_tac [("M","M")] Multiset.induction 1);
    64   by (simp_tac (arith_ss addsimps 
    64   by (simp_tac (arith_ss addsimps 
    66                            Multiset.countm_empty_def]) 1);
    66                            Multiset.countm_empty_def]) 1);
    67   by (asm_simp_tac (arith_ss addsimps 
    67   by (asm_simp_tac (arith_ss addsimps 
    68                        [Multiset.count_def,Multiset.delm_nonempty_def,
    68                        [Multiset.count_def,Multiset.delm_nonempty_def,
    69                         Multiset.countm_nonempty_def]
    69                         Multiset.countm_nonempty_def]
    70                     setloop (split_tac [expand_if])) 1);
    70                     setloop (split_tac [expand_if])) 1);
    71 val pos_count_imp_pos_countm = standard(result() RS mp);
    71 val pos_count_imp_pos_countm = store_thm("pos_count_imp_pos_countm", standard(result() RS mp));
    72 
    72 
    73 goal Multiset.thy
    73 goal Multiset.thy
    74    "!!P. P(x) ==> 0<count(M,x) --> countm(delm(M,x),P) = pred(countm(M,P))";
    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);
    75   by (res_inst_tac [("M","M")] Multiset.induction 1);
    76   by (simp_tac (arith_ss addsimps 
    76   by (simp_tac (arith_ss addsimps 
    79   by (asm_simp_tac (arith_ss addsimps 
    79   by (asm_simp_tac (arith_ss addsimps 
    80                       [eq_sym_conv,count_addm_simp,Multiset.delm_nonempty_def,
    80                       [eq_sym_conv,count_addm_simp,Multiset.delm_nonempty_def,
    81                        Multiset.countm_nonempty_def,pos_count_imp_pos_countm,
    81                        Multiset.countm_nonempty_def,pos_count_imp_pos_countm,
    82                        suc_pred_id]
    82                        suc_pred_id]
    83                     setloop (split_tac [expand_if])) 1);
    83                     setloop (split_tac [expand_if])) 1);
    84 val countm_done_delm = result();
    84 qed "countm_done_delm";