src/HOL/UNITY/AllocBase.thy
author paulson
Fri, 26 May 2000 18:08:46 +0200
changeset 8989 8791e3304748
parent 8974 a76f80911eb9
child 9024 b1f37f6819c4
permissions -rw-r--r--
sublist and some lemmas about it
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
Common declarations for Chandy and Charpentier's Allocator
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
     7
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
     8
with_path "../Induct" time_use_thy "AllocBase";
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
     9
*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    10
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
    11
AllocBase = Rename + Follows + MultisetOrder +
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    12
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    13
consts
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    14
  NbT      :: nat       (*Number of tokens in system*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    15
  Nclients :: nat       (*Number of clients*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    16
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    17
rules
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    18
  NbT_pos  "0 < NbT"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    19
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    20
(*This function merely sums the elements of a list*)
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    21
consts tokens     :: nat list => nat
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    22
primrec 
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    23
  "tokens [] = 0"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    24
  "tokens (x#xs) = x + tokens xs"
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    25
8974
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    26
(*Or could be setsum...(lessThan n)*)
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    27
consts sum_below :: "[nat=>'a, nat] => ('a::plus_ac0)"
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    28
primrec 
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    29
  sum_below_0    "sum_below f 0 = 0"
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    30
  sum_below_Suc  "sum_below f (Suc n) = f(n) + sum_below f n"
a76f80911eb9 sum_below moved here from Arith
paulson
parents: 8928
diff changeset
    31
8989
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
    32
constdefs sublist :: "['a list, nat set] => 'a list"
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
    33
    "sublist l A == map fst (filter (%p. snd p : A) (zip l [0..size l(]))"
8791e3304748 sublist and some lemmas about it
paulson
parents: 8974
diff changeset
    34
8928
1d3bf47a4ecc new files for the Allocator
paulson
parents:
diff changeset
    35
end