src/HOL/UNITY/Comp/AllocBase.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 65956 639eb3617a86
child 67613 ce654b0e6d69
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
     1
(*  Title:      HOL/UNITY/Comp/AllocBase.thy
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
*)
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 62608
diff changeset
     6
section\<open>Common Declarations for Chandy and Charpentier's Allocator\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
     7
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 65956
diff changeset
     8
theory AllocBase imports "../UNITY_Main" "HOL-Library.Multiset_Order" begin
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     9
45827
66c68453455c modernized specifications;
wenzelm
parents: 39246
diff changeset
    10
consts Nclients :: nat  (*Number of clients*)
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    11
45827
66c68453455c modernized specifications;
wenzelm
parents: 39246
diff changeset
    12
axiomatization NbT :: nat  (*Number of tokens in system*)
66c68453455c modernized specifications;
wenzelm
parents: 39246
diff changeset
    13
  where NbT_pos: "0 < NbT"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    14
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    15
abbreviation (input) tokens :: "nat list \<Rightarrow> nat"
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    16
where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63410
diff changeset
    17
  "tokens \<equiv> sum_list"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    18
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    19
abbreviation (input)
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    20
  "bag_of \<equiv> mset"
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    21
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    22
lemma sum_fun_mono:
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    23
  fixes f :: "nat \<Rightarrow> nat"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    24
  shows "(\<And>i. i < n \<Longrightarrow> f i \<le> g i) \<Longrightarrow> sum f {..<n} \<le> sum g {..<n}"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    25
  by (induct n) (auto simp add: lessThan_Suc add_le_mono)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    26
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    27
lemma tokens_mono_prefix: "xs \<le> ys \<Longrightarrow> tokens xs \<le> tokens ys"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 45827
diff changeset
    28
  by (induct ys arbitrary: xs) (auto simp add: prefix_Cons)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    29
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    30
lemma mono_tokens: "mono tokens"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    31
  using tokens_mono_prefix by (rule monoI)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    32
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    33
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    34
(** bag_of **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    35
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    36
lemma bag_of_append [simp]: "bag_of (l@l') = bag_of l + bag_of l'"
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    37
  by (fact mset_append)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    38
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    39
lemma mono_bag_of: "mono (bag_of :: 'a list => ('a::order) multiset)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    40
apply (rule monoI)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    41
apply (unfold prefix_def)
60397
f8a513fedb31 Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 58889
diff changeset
    42
apply (erule genPrefix.induct, simp_all add: add_right_mono)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    43
apply (erule order_trans)
63410
9789ccc2a477 more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 63409
diff changeset
    44
apply simp
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    45
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    46
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    47
(** sum **)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    48
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    49
declare sum.cong [cong]
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    50
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    51
lemma bag_of_nths_lemma:
15074
277b3a4da341 Modified \<Sum> syntax a little.
nipkow
parents: 15045
diff changeset
    52
     "(\<Sum>i\<in> A Int lessThan k. {#if i<k then f i else g i#}) =  
277b3a4da341 Modified \<Sum> syntax a little.
nipkow
parents: 15045
diff changeset
    53
      (\<Sum>i\<in> A Int lessThan k. {#f i#})"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    54
by (rule sum.cong, auto)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    55
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    56
lemma bag_of_nths:
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    57
     "bag_of (nths l A) =  
15074
277b3a4da341 Modified \<Sum> syntax a little.
nipkow
parents: 15045
diff changeset
    58
      (\<Sum>i\<in> A Int lessThan (length l). {# l!i #})"
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    59
  by (rule_tac xs = l in rev_induct)
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    60
     (simp_all add: nths_append Int_insert_right lessThan_Suc nth_append 
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    61
                    bag_of_nths_lemma ac_simps)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    62
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    63
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    64
lemma bag_of_nths_Un_Int:
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    65
     "bag_of (nths l (A Un B)) + bag_of (nths l (A Int B)) =  
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    66
      bag_of (nths l A) + bag_of (nths l B)"
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14738
diff changeset
    67
apply (subgoal_tac "A Int B Int {..<length l} =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24147
diff changeset
    68
                    (A Int {..<length l}) Int (B Int {..<length l}) ")
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    69
apply (simp add: bag_of_nths Int_Un_distrib2 sum.union_inter, blast)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    70
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    71
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    72
lemma bag_of_nths_Un_disjoint:
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    73
     "A Int B = {}  
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    74
      ==> bag_of (nths l (A Un B)) =  
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    75
          bag_of (nths l A) + bag_of (nths l B)"
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    76
by (simp add: bag_of_nths_Un_Int [symmetric])
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    77
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    78
lemma bag_of_nths_UN_disjoint [rule_format]:
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    79
     "[| finite I; ALL i:I. ALL j:I. i~=j --> A i Int A j = {} |]  
65956
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    80
      ==> bag_of (nths l (UNION I A)) =   
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    81
          (\<Sum>i\<in>I. bag_of (nths l (A i)))"
639eb3617a86 reorganised material on sublists
eberlm <eberlm@in.tum.de>
parents: 64267
diff changeset
    82
apply (auto simp add: bag_of_nths)
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    83
unfolding UN_simps [symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    84
apply (subst sum.UNION_disjoint)
62608
19f87fa0cfcb more theorems on orderings
haftmann
parents: 62430
diff changeset
    85
apply auto
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    86
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13786
diff changeset
    87
11194
ea13ff5a26d1 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    88
end