| author | haftmann | 
| Thu, 16 Apr 2009 14:02:11 +0200 | |
| changeset 30934 | ed5377c2b0a3 | 
| parent 18556 | dc39832e9280 | 
| 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/Counter | 
| 
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: Sidi O Ehmety, Cambridge University Computer Laboratory | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 4 | Copyright 2001 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 | From Charpentier and Chandy, | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 7 | Examples of Program Composition Illustrating the Use of Universal Properties | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 8 | In J. Rolim (editor), Parallel and Distributed Processing, | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 9 | Spriner LNCS 1586 (1999), pages 1215-1227. | 
| 
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 | |
| 14094 | 12 | header{*A Family of Similar Counters: Original Version*}
 | 
| 13 | ||
| 18556 | 14 | theory Counter imports "../UNITY_Main" begin | 
| 13792 | 15 | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 16 | (* Variables are names *) | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 17 | datatype name = C | c nat | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 18 | types state = "name=>int" | 
| 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 | consts | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 21 | sum :: "[nat,state]=>int" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 22 | sumj :: "[nat, nat, state]=>int" | 
| 
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 | primrec (* sum I s = sigma_{i<I}. s (c i) *)
 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 25 | "sum 0 s = 0" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 26 | "sum (Suc i) s = s (c i) + sum i s" | 
| 
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 | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 29 | "sumj 0 i s = 0" | 
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 30 | "sumj (Suc n) i s = (if n=i then sum n s else s (c n) + sumj n i s)" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 31 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 32 | types command = "(state*state)set" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 33 | |
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 34 | constdefs | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 35 | a :: "nat=>command" | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11701diff
changeset | 36 |  "a i == {(s, s'). s'=s(c i:= s (c i) + 1, C:= s C + 1)}"
 | 
| 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 | Component :: "nat => state program" | 
| 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 39 | "Component i == | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13792diff
changeset | 40 |     mk_total_program({s. s C = 0 & s (c i) = 0}, {a i},
 | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 41 | \<Union>G \<in> preserves (%s. s (c i)). Acts G)" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 42 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 43 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 44 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 45 | declare Component_def [THEN def_prg_Init, simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 46 | declare a_def [THEN def_act_simp, simp] | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 47 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 48 | (* Theorems about sum and sumj *) | 
| 14094 | 49 | lemma sum_upd_gt [rule_format]: "\<forall>n. I<n --> sum I (s(c n := x)) = sum I s" | 
| 50 | by (induct_tac "I", auto) | |
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 51 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 52 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 53 | lemma sum_upd_eq: "sum I (s(c I := x)) = sum I s" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 54 | apply (induct_tac "I") | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 55 | apply (auto simp add: sum_upd_gt [unfolded fun_upd_def]) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 56 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 57 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 58 | lemma sum_upd_C: "sum I (s(C := x)) = sum I s" | 
| 14094 | 59 | by (induct_tac "I", auto) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 60 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 61 | lemma sumj_upd_ci: "sumj I i (s(c i := x)) = sumj I i s" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 62 | apply (induct_tac "I") | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 63 | apply (auto simp add: sum_upd_eq [unfolded fun_upd_def]) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 64 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 65 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 66 | lemma sumj_upd_C: "sumj I i (s(C := x)) = sumj I i s" | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 67 | apply (induct_tac "I") | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 68 | apply (auto simp add: sum_upd_C [unfolded fun_upd_def]) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 69 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 70 | |
| 14094 | 71 | lemma sumj_sum_gt [rule_format]: "\<forall>i. I<i--> (sumj I i s = sum I s)" | 
| 72 | by (induct_tac "I", auto) | |
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 73 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 74 | lemma sumj_sum_eq: "(sumj I I s = sum I s)" | 
| 14094 | 75 | apply (induct_tac "I", auto) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 76 | apply (simp (no_asm) add: sumj_sum_gt) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 77 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 78 | |
| 14094 | 79 | lemma sum_sumj [rule_format]: "\<forall>i. i<I-->(sum I s = s (c i) + sumj I i s)" | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 80 | apply (induct_tac "I") | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 81 | apply (auto simp add: linorder_neq_iff sumj_sum_eq) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 82 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 83 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 84 | (* Correctness proofs for Components *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 85 | (* p2 and p3 proofs *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 86 | lemma p2: "Component i \<in> stable {s. s C = s (c i) + k}"
 | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14094diff
changeset | 87 | by (simp add: Component_def, safety) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 88 | |
| 14094 | 89 | lemma p3: "Component i \<in> stable {s. \<forall>v. v\<noteq>c i & v\<noteq>C --> s v = k v}"
 | 
| 16184 
80617b8d33c5
renamed "constrains" to "safety" to avoid keyword clash
 paulson parents: 
14094diff
changeset | 90 | by (simp add: Component_def, safety) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 91 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 92 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 93 | lemma p2_p3_lemma1: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 94 | "(\<forall>k. Component i \<in> stable ({s. s C = s (c i) + sumj I i k}  
 | 
| 14094 | 95 |                    \<inter> {s. \<forall>v. v\<noteq>c i & v\<noteq>C --> s v = k v}))  
 | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 96 |    = (Component i \<in> stable {s. s C = s (c i) + sumj I i s})"
 | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 97 | apply (simp add: Component_def mk_total_program_def) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 98 | apply (auto simp add: constrains_def stable_def sumj_upd_C sumj_upd_ci) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 99 | done | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 100 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 101 | lemma p2_p3_lemma2: | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 102 | "\<forall>k. Component i \<in> stable ({s. s C = s (c i) + sumj I i k} Int  
 | 
| 14094 | 103 |                             {s. \<forall>v. v\<noteq>c i & v\<noteq>C --> s v = k v})"
 | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 104 | by (blast intro: stable_Int [OF p2 p3]) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 105 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 106 | lemma p2_p3: "Component i \<in> stable {s.  s C = s (c i) + sumj I i s}"
 | 
| 14094 | 107 | by (auto intro!: p2_p3_lemma2 simp add: p2_p3_lemma1 [symmetric]) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 108 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 109 | (* Compositional Proof *) | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 110 | |
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 111 | lemma sum_0' [rule_format]: "(\<forall>i. i < I --> s (c i) = 0) --> sum I s = 0" | 
| 14094 | 112 | by (induct_tac "I", auto) | 
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 113 | |
| 14094 | 114 | (* I cannot be empty *) | 
| 115 | lemma safety: | |
| 116 |      "0<I ==> (\<Squnion>i \<in> {i. i<I}. Component i) \<in> invariant {s. s C = sum I s}"
 | |
| 117 | apply (simp (no_asm) add: invariant_def JN_stable sum_sumj) | |
| 14088 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 118 | apply (force intro: p2_p3 sum_0') | 
| 
61bd46feb919
converted Counter, Counterc and PriorityAux to Isar scripts (all HOL/UNITY/Comp)
 paulson parents: 
13812diff
changeset | 119 | done | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13792diff
changeset | 120 | |
| 11194 
ea13ff5a26d1
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 paulson parents: diff
changeset | 121 | end |