| author | wenzelm | 
| Thu, 26 Apr 2007 15:42:04 +0200 | |
| changeset 22814 | 4cd25f1706bb | 
| parent 16417 | 9bc16273c2d4 | 
| child 32960 | 69916a850301 | 
| 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/Client.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 | |
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 7 | header{*Distributed Resource Management System: the Client*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 8 | |
| 16417 | 9 | theory Client imports Rename AllocBase begin | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 10 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 11 | types | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 12 |   tokbag = nat	   --{*tokbags could be multisets...or any ordered type?*}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 13 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 14 | record state = | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 15 |   giv :: "tokbag list" --{*input history: tokens granted*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 16 |   ask :: "tokbag list" --{*output history: tokens requested*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 17 |   rel :: "tokbag list" --{*output history: tokens released*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 18 |   tok :: tokbag	       --{*current token request*}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 19 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 20 | record 'a state_d = | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | state + | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 22 |   dummy :: 'a          --{*new variables*}
 | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 23 | |
| 
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 | (*Array indexing is translated to list indexing as A[n] == A!(n-1). *) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 26 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 27 | constdefs | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 28 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 29 | (** Release some tokens **) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 30 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 31 |   rel_act :: "('a state_d * 'a state_d) set"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 32 |     "rel_act == {(s,s').
 | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 33 | \<exists>nrel. nrel = size (rel s) & | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
11194diff
changeset | 34 | s' = s (| rel := rel s @ [giv s!nrel] |) & | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
11194diff
changeset | 35 | nrel < size (giv s) & | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 36 | ask s!nrel \<le> giv s!nrel}" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 37 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 38 | (** Choose a new token requirement **) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 39 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 40 | (** Including s'=s suppresses fairness, allowing the non-trivial part | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 41 | of the action to be ignored **) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 42 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 43 |   tok_act :: "('a state_d * 'a state_d) set"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 44 |      "tok_act == {(s,s'). s'=s | s' = s (|tok := Suc (tok s mod NbT) |)}"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 45 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 46 |   ask_act :: "('a state_d * 'a state_d) set"
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 47 |     "ask_act == {(s,s'). s'=s |
 | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 48 | (s' = s (|ask := ask s @ [tok s]|))}" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 49 | |
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 50 | Client :: "'a state_d program" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 51 | "Client == | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
11194diff
changeset | 52 | mk_total_program | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 53 |             ({s. tok s \<in> atMost NbT &
 | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
11194diff
changeset | 54 | giv s = [] & ask s = [] & rel s = []}, | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
11194diff
changeset | 55 | 	     {rel_act, tok_act, ask_act},
 | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 56 | \<Union>G \<in> preserves rel Int preserves ask Int preserves tok. | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 57 | Acts G)" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 58 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 59 | (*Maybe want a special theory section to declare such maps*) | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 60 | non_dummy :: "'a state_d => state" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 61 | "non_dummy s == (|giv = giv s, ask = ask s, rel = rel s, tok = tok s|)" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 62 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 63 | (*Renaming map to put a Client into the standard form*) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 64 | client_map :: "'a state_d => state*'a" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 65 | "client_map == funPair non_dummy dummy" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 66 | |
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 67 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 68 | declare Client_def [THEN def_prg_Init, simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 69 | declare Client_def [THEN def_prg_AllowedActs, simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 70 | declare rel_act_def [THEN def_act_simp, simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 71 | declare tok_act_def [THEN def_act_simp, simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 72 | declare ask_act_def [THEN def_act_simp, simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 73 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 74 | lemma Client_ok_iff [iff]: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 75 | "(Client ok G) = | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 76 | (G \<in> preserves rel & G \<in> preserves ask & G \<in> preserves tok & | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 77 | Client \<in> Allowed G)" | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 78 | by (auto simp add: ok_iff_Allowed Client_def [THEN def_total_prg_Allowed]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 79 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 80 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 81 | text{*Safety property 1: ask, rel are increasing*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 82 | lemma increasing_ask_rel: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 83 | "Client \<in> UNIV guarantees Increasing ask Int Increasing rel" | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 84 | apply (auto intro!: increasing_imp_Increasing simp add: guar_def preserves_subset_increasing [THEN subsetD]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 85 | apply (auto simp add: Client_def increasing_def) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14089diff
changeset | 86 | apply (safety, auto)+ | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 87 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 88 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 89 | declare nth_append [simp] append_one_prefix [simp] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 90 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 91 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 92 | text{*Safety property 2: the client never requests too many tokens.
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 93 | With no Substitution Axiom, we must prove the two invariants | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 94 | simultaneously. *} | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 95 | lemma ask_bounded_lemma: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 96 | "Client ok G | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 97 | ==> Client Join G \<in> | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 98 |               Always ({s. tok s \<le> NbT}  Int   
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 99 |                       {s. \<forall>elt \<in> set (ask s). elt \<le> NbT})"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 100 | apply auto | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 101 | apply (rule invariantI [THEN stable_Join_Always2], force) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 102 | prefer 2 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 103 | apply (fast elim!: preserves_subset_stable [THEN subsetD] intro!: stable_Int) | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14089diff
changeset | 104 | apply (simp add: Client_def, safety) | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 105 | apply (cut_tac m = "tok s" in NbT_pos [THEN mod_less_divisor], auto) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 106 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 107 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 108 | text{*export version, with no mention of tok in the postcondition, but
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 109 | unfortunately tok must be declared local.*} | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 110 | lemma ask_bounded: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 111 |      "Client \<in> UNIV guarantees Always {s. \<forall>elt \<in> set (ask s). elt \<le> NbT}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 112 | apply (rule guaranteesI) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 113 | apply (erule ask_bounded_lemma [THEN Always_weaken]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 114 | apply (rule Int_lower2) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 115 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 116 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 117 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 118 | text{*** Towards proving the liveness property ***}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 119 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 120 | lemma stable_rel_le_giv: "Client \<in> stable {s. rel s \<le> giv s}"
 | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14089diff
changeset | 121 | by (simp add: Client_def, safety, auto) | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 122 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 123 | lemma Join_Stable_rel_le_giv: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 124 | "[| Client Join G \<in> Increasing giv; G \<in> preserves rel |] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 125 |       ==> Client Join G \<in> Stable {s. rel s \<le> giv s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 126 | by (rule stable_rel_le_giv [THEN Increasing_preserves_Stable], auto) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 127 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 128 | lemma Join_Always_rel_le_giv: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 129 | "[| Client Join G \<in> Increasing giv; G \<in> preserves rel |] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 130 |       ==> Client Join G \<in> Always {s. rel s \<le> giv s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 131 | by (force intro: AlwaysI Join_Stable_rel_le_giv) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 132 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 133 | lemma transient_lemma: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 134 |      "Client \<in> transient {s. rel s = k & k<h & h \<le> giv s & h pfixGe ask s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 135 | apply (simp add: Client_def mk_total_program_def) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 136 | apply (rule_tac act = rel_act in totalize_transientI) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 137 | apply (auto simp add: Domain_def Client_def) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 138 | apply (blast intro: less_le_trans prefix_length_le strict_prefix_length_less) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 139 | apply (auto simp add: prefix_def genPrefix_iff_nth Ge_def) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 140 | apply (blast intro: strict_prefix_length_less) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 141 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 142 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 143 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 144 | lemma induct_lemma: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 145 | "[| Client Join G \<in> Increasing giv; Client ok G |] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 146 |   ==> Client Join G \<in> {s. rel s = k & k<h & h \<le> giv s & h pfixGe ask s}   
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 147 |                       LeadsTo {s. k < rel s & rel s \<le> giv s &  
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 148 | h \<le> giv s & h pfixGe ask s}" | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 149 | apply (rule single_LeadsTo_I) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 150 | apply (frule increasing_ask_rel [THEN guaranteesD], auto) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 151 | apply (rule transient_lemma [THEN Join_transient_I1, THEN transient_imp_leadsTo, THEN leadsTo_imp_LeadsTo, THEN PSP_Stable, THEN LeadsTo_weaken]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 152 | apply (rule Stable_Int [THEN Stable_Int, THEN Stable_Int]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 153 | apply (erule_tac f = giv and x = "giv s" in IncreasingD) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 154 | apply (erule_tac f = ask and x = "ask s" in IncreasingD) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 155 | apply (erule_tac f = rel and x = "rel s" in IncreasingD) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 156 | apply (erule Join_Stable_rel_le_giv, blast) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 157 | apply (blast intro: order_less_imp_le order_trans) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 158 | apply (blast intro: sym order_less_le [THEN iffD2] order_trans | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 159 | prefix_imp_pfixGe pfixGe_trans) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 160 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 161 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 162 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 163 | lemma rel_progress_lemma: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 164 | "[| Client Join G \<in> Increasing giv; Client ok G |] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 165 |   ==> Client Join G \<in> {s. rel s < h & h \<le> giv s & h pfixGe ask s}   
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 166 |                       LeadsTo {s. h \<le> rel s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 167 | apply (rule_tac f = "%s. size h - size (rel s) " in LessThan_induct) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 168 | apply (auto simp add: vimage_def) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 169 | apply (rule single_LeadsTo_I) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 170 | apply (rule induct_lemma [THEN LeadsTo_weaken], auto) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 171 | apply (blast intro: order_less_le [THEN iffD2] dest: common_prefix_linear) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 172 | apply (drule strict_prefix_length_less)+ | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 173 | apply arith | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 174 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 175 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 176 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 177 | lemma client_progress_lemma: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 178 | "[| Client Join G \<in> Increasing giv; Client ok G |] | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 179 |       ==> Client Join G \<in> {s. h \<le> giv s & h pfixGe ask s}   
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 180 |                           LeadsTo {s. h \<le> rel s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 181 | apply (rule Join_Always_rel_le_giv [THEN Always_LeadsToI], simp_all) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 182 | apply (rule LeadsTo_Un [THEN LeadsTo_weaken_L]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 183 | apply (blast intro: rel_progress_lemma) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 184 | apply (rule subset_refl [THEN subset_imp_LeadsTo]) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 185 | apply (blast intro: order_less_le [THEN iffD2] dest: common_prefix_linear) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 186 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 187 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 188 | text{*Progress property: all tokens that are given will be released*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 189 | lemma client_progress: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 190 | "Client \<in> | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 191 | Increasing giv guarantees | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 192 |        (INT h. {s. h \<le> giv s & h pfixGe ask s} LeadsTo {s. h \<le> rel s})"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 193 | apply (rule guaranteesI, clarify) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 194 | apply (blast intro: client_progress_lemma) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 195 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 196 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 197 | text{*This shows that the Client won't alter other variables in any state
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 198 | that it is combined with*} | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 199 | lemma client_preserves_dummy: "Client \<in> preserves dummy" | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14089diff
changeset | 200 | by (simp add: Client_def preserves_def, clarify, safety, auto) | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 201 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 202 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 203 | text{** Obsolete lemmas from first version of the Client **}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 204 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 205 | lemma stable_size_rel_le_giv: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 206 |      "Client \<in> stable {s. size (rel s) \<le> size (giv s)}"
 | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14089diff
changeset | 207 | by (simp add: Client_def, safety, auto) | 
| 14089 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 208 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 209 | text{*clients return the right number of tokens*}
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 210 | lemma ok_guar_rel_prefix_giv: | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 211 |      "Client \<in> Increasing giv  guarantees  Always {s. rel s \<le> giv s}"
 | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 212 | apply (rule guaranteesI) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 213 | apply (rule AlwaysI, force) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 214 | apply (blast intro: Increasing_preserves_Stable stable_rel_le_giv) | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 215 | done | 
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 216 | |
| 
7b34f58b1b81
converted UNITY/Comp/{AllocImpl,Client} to Isar scripts
 paulson parents: 
13812diff
changeset | 217 | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 218 | end |