src/HOL/UNITY/Comp/AllocBase.thy
changeset 66453 cc19f7ca2ed6
parent 65956 639eb3617a86
child 67613 ce654b0e6d69
equal deleted inserted replaced
66452:450cefec7c11 66453:cc19f7ca2ed6
     3     Copyright   1998  University of Cambridge
     3     Copyright   1998  University of Cambridge
     4 *)
     4 *)
     5 
     5 
     6 section\<open>Common Declarations for Chandy and Charpentier's Allocator\<close>
     6 section\<open>Common Declarations for Chandy and Charpentier's Allocator\<close>
     7 
     7 
     8 theory AllocBase imports "../UNITY_Main" "~~/src/HOL/Library/Multiset_Order" begin
     8 theory AllocBase imports "../UNITY_Main" "HOL-Library.Multiset_Order" begin
     9 
     9 
    10 consts Nclients :: nat  (*Number of clients*)
    10 consts Nclients :: nat  (*Number of clients*)
    11 
    11 
    12 axiomatization NbT :: nat  (*Number of tokens in system*)
    12 axiomatization NbT :: nat  (*Number of tokens in system*)
    13   where NbT_pos: "0 < NbT"
    13   where NbT_pos: "0 < NbT"