src/HOL/UNITY/AllocBase.ML
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 9747 043098ba5098
child 10265 4e004b548049
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/AllocBase
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     2
    ID:         $Id$
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     5
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     6
Basis declarations for Chandy and Charpentier's Allocator
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8928
diff changeset
     7
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
     8
add_path "../Induct";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
     9
time_use_thy "AllocBase";
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    10
*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    11
9109
0085c32a533b sum_below f n -> setsum f (lessThan n)
paulson
parents: 9107
diff changeset
    12
Goal "!!f :: nat=>nat. \
0085c32a533b sum_below f n -> setsum f (lessThan n)
paulson
parents: 9107
diff changeset
    13
\     (ALL i. i<n --> f i <= g i) --> \
0085c32a533b sum_below f n -> setsum f (lessThan n)
paulson
parents: 9107
diff changeset
    14
\     setsum f (lessThan n) <= setsum g (lessThan n)";
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    15
by (induct_tac "n" 1);
9109
0085c32a533b sum_below f n -> setsum f (lessThan n)
paulson
parents: 9107
diff changeset
    16
by (auto_tac (claset(), simpset() addsimps [lessThan_Suc]));
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    17
by (dres_inst_tac [("x","n")] spec 1);
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    18
by (arith_tac 1);
9109
0085c32a533b sum_below f n -> setsum f (lessThan n)
paulson
parents: 9107
diff changeset
    19
qed_spec_mp "setsum_fun_mono";
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    20
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    21
Goal "ALL xs. xs <= ys --> tokens xs <= tokens ys";
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    22
by (induct_tac "ys" 1);
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    23
by (auto_tac (claset(), simpset() addsimps [prefix_Cons]));
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    24
qed_spec_mp "tokens_mono_prefix";
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    25
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    26
Goalw [mono_def] "mono tokens";
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    27
by (blast_tac (claset() addIs [tokens_mono_prefix]) 1);
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    28
qed "mono_tokens";
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8928
diff changeset
    29
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    30
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    31
(** bag_of **)
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8928
diff changeset
    32
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    33
Goal "bag_of (l@l') = bag_of l + bag_of l'";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    34
by (induct_tac "l" 1);
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    35
 by (asm_simp_tac (simpset() addsimps (thms "plus_ac0")) 2);
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    36
by (Simp_tac 1);
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    37
qed "bag_of_append";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    38
Addsimps [bag_of_append];
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    39
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    40
Goal "mono (bag_of :: 'a list => ('a::order) multiset)";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    41
by (rtac monoI 1); 
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    42
by (rewtac prefix_def);
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    43
by (etac genPrefix.induct 1);
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    44
by Auto_tac;
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    45
by (asm_full_simp_tac (simpset() addsimps [union_le_mono]) 1); 
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    46
by (etac order_trans 1); 
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    47
by (rtac union_upper1 1); 
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    48
qed "mono_bag_of";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    49
9107
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    50
(** setsum **)
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    51
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    52
Addcongs [setsum_cong];
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    53
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    54
Goal "setsum (%i. {#if i<k then f i else g i#}) (A Int lessThan k) = \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    55
\     setsum (%i. {#f i#}) (A Int lessThan k)";
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    56
by (rtac setsum_cong 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    57
by Auto_tac;  
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    58
qed "bag_of_sublist_lemma";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    59
9107
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    60
Goal "bag_of (sublist l A) = \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    61
\     setsum (%i. {# l!i #}) (A Int lessThan (length l))";
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9336
diff changeset
    62
by (rev_induct_tac "l" 1);
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    63
by (Simp_tac 1);
9107
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    64
by (asm_simp_tac
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    65
    (simpset() addsimps [sublist_append, Int_insert_right, lessThan_Suc, 
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    66
                    nth_append, bag_of_sublist_lemma] @ thms "plus_ac0") 1);
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    67
qed "bag_of_sublist";
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    68
9107
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    69
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    70
Goal "bag_of (sublist l (A Un B)) + bag_of (sublist l (A Int B)) = \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    71
\     bag_of (sublist l A) + bag_of (sublist l B)";
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    72
by (subgoal_tac "A Int B Int {..length l(} = \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    73
\                (A Int {..length l(}) Int (B Int {..length l(})" 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    74
by (asm_simp_tac (simpset() addsimps [bag_of_sublist, Int_Un_distrib2, 
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    75
                                      setsum_Un_Int]) 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    76
by (Blast_tac 1);
9026
640c4fd8b5b3 lots of new results about sublist, bag_of
paulson
parents: 8989
diff changeset
    77
qed "bag_of_sublist_Un_Int";
9107
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    78
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    79
Goal "A Int B = {} \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    80
\     ==> bag_of (sublist l (A Un B)) = \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    81
\         bag_of (sublist l A) + bag_of (sublist l B)"; 
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    82
by (asm_simp_tac (simpset() addsimps [bag_of_sublist_Un_Int RS sym]) 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    83
qed "bag_of_sublist_Un_disjoint";
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    84
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    85
Goal "[| finite I; ALL i:I. ALL j:I. i~=j --> A i Int A j = {} |] \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    86
\     ==> bag_of (sublist l (UNION I A)) =  \
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    87
\         setsum (%i. bag_of (sublist l (A i))) I";  
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    88
by (asm_simp_tac (simpset() delsimps UN_simps addsimps (UN_simps RL [sym])
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    89
			    addsimps [bag_of_sublist]) 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    90
by (stac setsum_UN_disjoint 1);
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    91
by Auto_tac;  
67202a50ee6d removed some finiteness conditions from bag_of/sublist theorems
paulson
parents: 9090
diff changeset
    92
qed_spec_mp "bag_of_sublist_UN_disjoint";