| author | mengj | 
| Tue, 07 Mar 2006 04:04:21 +0100 | |
| changeset 19200 | 1da6b9a1575d | 
| parent 18556 | dc39832e9280 | 
| child 20217 | 25b068a99d2b | 
| permissions | -rw-r--r-- | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/AllocBase.thy | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 1998 University of Cambridge | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 5 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 6 | *) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | |
| 13798 | 8 | header{*Common Declarations for Chandy and Charpentier's Allocator*}
 | 
| 9 | ||
| 18556 | 10 | theory AllocBase imports "../UNITY_Main" begin | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 12 | consts | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | NbT :: nat (*Number of tokens in system*) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | Nclients :: nat (*Number of clients*) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 15 | |
| 13798 | 16 | axioms | 
| 17 | NbT_pos: "0 < NbT" | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 18 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 19 | (*This function merely sums the elements of a list*) | 
| 13798 | 20 | consts tokens :: "nat list => nat" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | primrec | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | "tokens [] = 0" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | "tokens (x#xs) = x + tokens xs" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 24 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 25 | consts | 
| 13798 | 26 | bag_of :: "'a list => 'a multiset" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 27 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 28 | primrec | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 |   "bag_of []     = {#}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 30 |   "bag_of (x#xs) = {#x#} + bag_of xs"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 31 | |
| 13798 | 32 | lemma setsum_fun_mono [rule_format]: | 
| 33 | "!!f :: nat=>nat. | |
| 34 | (ALL i. i<n --> f i <= g i) --> | |
| 35 | setsum f (lessThan n) <= setsum g (lessThan n)" | |
| 36 | apply (induct_tac "n") | |
| 37 | apply (auto simp add: lessThan_Suc) | |
| 38 | apply (drule_tac x = n in spec, arith) | |
| 39 | done | |
| 40 | ||
| 41 | lemma tokens_mono_prefix [rule_format]: | |
| 42 | "ALL xs. xs <= ys --> tokens xs <= tokens ys" | |
| 43 | apply (induct_tac "ys") | |
| 44 | apply (auto simp add: prefix_Cons) | |
| 45 | done | |
| 46 | ||
| 47 | lemma mono_tokens: "mono tokens" | |
| 48 | apply (unfold mono_def) | |
| 49 | apply (blast intro: tokens_mono_prefix) | |
| 50 | done | |
| 51 | ||
| 52 | ||
| 53 | (** bag_of **) | |
| 54 | ||
| 55 | lemma bag_of_append [simp]: "bag_of (l@l') = bag_of l + bag_of l'" | |
| 56 | apply (induct_tac "l", simp) | |
| 14738 | 57 | apply (simp add: add_ac) | 
| 13798 | 58 | done | 
| 59 | ||
| 60 | lemma mono_bag_of: "mono (bag_of :: 'a list => ('a::order) multiset)"
 | |
| 61 | apply (rule monoI) | |
| 62 | apply (unfold prefix_def) | |
| 63 | apply (erule genPrefix.induct, auto) | |
| 64 | apply (simp add: union_le_mono) | |
| 65 | apply (erule order_trans) | |
| 66 | apply (rule union_upper1) | |
| 67 | done | |
| 68 | ||
| 69 | (** setsum **) | |
| 70 | ||
| 71 | declare setsum_cong [cong] | |
| 72 | ||
| 73 | lemma bag_of_sublist_lemma: | |
| 15074 | 74 |      "(\<Sum>i\<in> A Int lessThan k. {#if i<k then f i else g i#}) =  
 | 
| 75 |       (\<Sum>i\<in> A Int lessThan k. {#f i#})"
 | |
| 14114 | 76 | by (rule setsum_cong, auto) | 
| 13798 | 77 | |
| 78 | lemma bag_of_sublist: | |
| 79 | "bag_of (sublist l A) = | |
| 15074 | 80 |       (\<Sum>i\<in> A Int lessThan (length l). {# l!i #})"
 | 
| 13798 | 81 | apply (rule_tac xs = l in rev_induct, simp) | 
| 82 | apply (simp add: sublist_append Int_insert_right lessThan_Suc nth_append | |
| 14738 | 83 | bag_of_sublist_lemma add_ac) | 
| 13798 | 84 | done | 
| 85 | ||
| 86 | ||
| 87 | lemma bag_of_sublist_Un_Int: | |
| 88 | "bag_of (sublist l (A Un B)) + bag_of (sublist l (A Int B)) = | |
| 89 | bag_of (sublist l A) + bag_of (sublist l B)" | |
| 15045 | 90 | apply (subgoal_tac "A Int B Int {..<length l} =
 | 
| 91 | 	            (A Int {..<length l}) Int (B Int {..<length l}) ")
 | |
| 13798 | 92 | apply (simp add: bag_of_sublist Int_Un_distrib2 setsum_Un_Int, blast) | 
| 93 | done | |
| 94 | ||
| 95 | lemma bag_of_sublist_Un_disjoint: | |
| 96 |      "A Int B = {}  
 | |
| 97 | ==> bag_of (sublist l (A Un B)) = | |
| 98 | bag_of (sublist l A) + bag_of (sublist l B)" | |
| 14114 | 99 | by (simp add: bag_of_sublist_Un_Int [symmetric]) | 
| 13798 | 100 | |
| 101 | lemma bag_of_sublist_UN_disjoint [rule_format]: | |
| 102 |      "[| finite I; ALL i:I. ALL j:I. i~=j --> A i Int A j = {} |]  
 | |
| 103 | ==> bag_of (sublist l (UNION I A)) = | |
| 15074 | 104 | (\<Sum>i\<in>I. bag_of (sublist l (A i)))" | 
| 13798 | 105 | apply (simp del: UN_simps | 
| 106 | add: UN_simps [symmetric] add: bag_of_sublist) | |
| 107 | apply (subst setsum_UN_disjoint, auto) | |
| 108 | done | |
| 109 | ||
| 110 | ML | |
| 111 | {*
 | |
| 112 | val NbT_pos = thm "NbT_pos"; | |
| 113 | val setsum_fun_mono = thm "setsum_fun_mono"; | |
| 114 | val tokens_mono_prefix = thm "tokens_mono_prefix"; | |
| 115 | val mono_tokens = thm "mono_tokens"; | |
| 116 | val bag_of_append = thm "bag_of_append"; | |
| 117 | val mono_bag_of = thm "mono_bag_of"; | |
| 118 | val bag_of_sublist_lemma = thm "bag_of_sublist_lemma"; | |
| 119 | val bag_of_sublist = thm "bag_of_sublist"; | |
| 120 | val bag_of_sublist_Un_Int = thm "bag_of_sublist_Un_Int"; | |
| 121 | val bag_of_sublist_Un_disjoint = thm "bag_of_sublist_Un_disjoint"; | |
| 122 | val bag_of_sublist_UN_disjoint = thm "bag_of_sublist_UN_disjoint"; | |
| 123 | *} | |
| 124 | ||
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 125 | end |