| author | nipkow | 
| Fri, 28 Aug 2009 20:21:50 +0200 | |
| changeset 32444 | bd13b7756f47 | 
| parent 21026 | 3b2821e0d541 | 
| child 32456 | 341c83339aeb | 
| permissions | -rw-r--r-- | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/UNITY/Common  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
4  | 
Copyright 1998 University of Cambridge  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
6  | 
Common Meeting Time example from Misra (1994)  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
8  | 
The state is identified with the one variable in existence.  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
10  | 
From Misra, "A Logic for Concurrent Programming" (1994), sections 5.1 and 13.1.  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
11  | 
*)  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
12  | 
|
| 18556 | 13  | 
theory Common imports "../UNITY_Main" begin  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
15  | 
consts  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
16  | 
ftime :: "nat=>nat"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
17  | 
gtime :: "nat=>nat"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
18  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
19  | 
axioms  | 
| 13806 | 20  | 
fmono: "m \<le> n ==> ftime m \<le> ftime n"  | 
21  | 
gmono: "m \<le> n ==> gtime m \<le> gtime n"  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
22  | 
|
| 13806 | 23  | 
fasc: "m \<le> ftime n"  | 
24  | 
gasc: "m \<le> gtime n"  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
26  | 
constdefs  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
27  | 
common :: "nat set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
28  | 
    "common == {n. ftime n = n & gtime n = n}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
29  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
30  | 
maxfg :: "nat => nat set"  | 
| 13806 | 31  | 
    "maxfg m == {t. t \<le> max (ftime m) (gtime m)}"
 | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
32  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
33  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
34  | 
(*Misra's property CMT4: t exceeds no common meeting time*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
35  | 
lemma common_stable:  | 
| 13806 | 36  | 
     "[| \<forall>m. F \<in> {m} Co (maxfg m); n \<in> common |]  
 | 
37  | 
==> F \<in> Stable (atMost n)"  | 
|
38  | 
apply (drule_tac M = "{t. t \<le> n}" in Elimination_sing)
 | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
39  | 
apply (simp add: atMost_def Stable_def common_def maxfg_def le_max_iff_disj)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
40  | 
apply (erule Constrains_weaken_R)  | 
| 21026 | 41  | 
apply (blast intro: order_eq_refl le_trans dest: fmono gmono)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
42  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
43  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
44  | 
lemma common_safety:  | 
| 13806 | 45  | 
"[| Init F \<subseteq> atMost n;  | 
46  | 
         \<forall>m. F \<in> {m} Co (maxfg m); n \<in> common |]  
 | 
|
47  | 
==> F \<in> Always (atMost n)"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
48  | 
by (simp add: AlwaysI common_stable)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
49  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
50  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
51  | 
(*** Some programs that implement the safety property above ***)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
52  | 
|
| 13806 | 53  | 
lemma "SKIP \<in> {m} co (maxfg m)"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
54  | 
by (simp add: constrains_def maxfg_def le_max_iff_disj fasc)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
55  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
56  | 
(*This one is t := ftime t || t := gtime t*)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
57  | 
lemma "mk_total_program  | 
| 
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
58  | 
         (UNIV, {range(%t.(t,ftime t)), range(%t.(t,gtime t))}, UNIV)
 | 
| 13806 | 59  | 
       \<in> {m} co (maxfg m)"
 | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
60  | 
apply (simp add: mk_total_program_def)  | 
| 
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
61  | 
apply (simp add: constrains_def maxfg_def le_max_iff_disj fasc)  | 
| 
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
62  | 
done  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
63  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
64  | 
(*This one is t := max (ftime t) (gtime t)*)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
65  | 
lemma "mk_total_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}, UNIV)
 | 
| 13806 | 66  | 
       \<in> {m} co (maxfg m)"
 | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
67  | 
apply (simp add: mk_total_program_def)  | 
| 32444 | 68  | 
apply (simp add: constrains_def maxfg_def gasc)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
69  | 
done  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
70  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
71  | 
(*This one is t := t+1 if t <max (ftime t) (gtime t) *)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
72  | 
lemma "mk_total_program  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
73  | 
          (UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }, UNIV)   
 | 
| 13806 | 74  | 
       \<in> {m} co (maxfg m)"
 | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
75  | 
apply (simp add: mk_total_program_def)  | 
| 32444 | 76  | 
apply (simp add: constrains_def maxfg_def gasc)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13806 
diff
changeset
 | 
77  | 
done  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
78  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
79  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
80  | 
(*It remans to prove that they satisfy CMT3': t does not decrease,  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
81  | 
and that CMT3' implies that t stops changing once common(t) holds.*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
82  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
83  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
84  | 
(*** Progress under weak fairness ***)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
85  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
86  | 
declare atMost_Int_atLeast [simp]  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
87  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
88  | 
lemma leadsTo_common_lemma:  | 
| 13806 | 89  | 
     "[| \<forall>m. F \<in> {m} Co (maxfg m);  
 | 
90  | 
         \<forall>m \<in> lessThan n. F \<in> {m} LeadsTo (greaterThan m);  
 | 
|
91  | 
n \<in> common |]  | 
|
92  | 
==> F \<in> (atMost n) LeadsTo common"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
93  | 
apply (rule LeadsTo_weaken_R)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
94  | 
apply (rule_tac f = id and l = n in GreaterThan_bounded_induct)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
95  | 
apply (simp_all (no_asm_simp))  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
96  | 
apply (rule_tac [2] subset_refl)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
97  | 
apply (blast dest: PSP_Stable2 intro: common_stable LeadsTo_weaken_R)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
98  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
99  | 
|
| 13806 | 100  | 
(*The "\<forall>m \<in> -common" form echoes CMT6.*)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
101  | 
lemma leadsTo_common:  | 
| 13806 | 102  | 
     "[| \<forall>m. F \<in> {m} Co (maxfg m);  
 | 
103  | 
         \<forall>m \<in> -common. F \<in> {m} LeadsTo (greaterThan m);  
 | 
|
104  | 
n \<in> common |]  | 
|
105  | 
==> F \<in> (atMost (LEAST n. n \<in> common)) LeadsTo common"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
106  | 
apply (rule leadsTo_common_lemma)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
107  | 
apply (simp_all (no_asm_simp))  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
108  | 
apply (erule_tac [2] LeastI)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
109  | 
apply (blast dest!: not_less_Least)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
110  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
111  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
112  | 
end  |