src/HOL/UNITY/Deadlock.ML
author paulson
Thu, 26 Aug 1999 11:33:24 +0200
changeset 7359 98a2afab3f86
parent 7127 48e235179ffb
child 8334 7896bcbd8641
permissions -rw-r--r--
extra syntax for JN, making it more like UN
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5111
diff changeset
     1
(*  Title:      HOL/UNITY/Deadlock
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     2
    ID:         $Id$
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     4
    Copyright   1998  University of Cambridge
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     6
Deadlock examples from section 5.6 of 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     7
    Misra, "A Logic for Concurrent Programming", 1994
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     8
*)
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
(*Trivial, two-process case*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    11
Goalw [constrains_def, stable_def]
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    12
    "[| F : (A Int B) co A;  F : (B Int A) co B |] ==> F : stable (A Int B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
(*a simplification step*)
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    18
Goal "(INT i: atMost n. A(Suc i) Int A i) = (INT i: atMost (Suc n). A i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
by (induct_tac "n" 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    20
by (ALLGOALS (asm_simp_tac (simpset() addsimps [atMost_Suc])));
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
by (blast_tac (claset() addEs [le_SucE] addSEs [equalityE]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
qed "Collect_le_Int_equals";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
(*Dual of the required property.  Converse inclusion fails.*)
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    26
Goal "(UN i: lessThan n. A i) Int (- A n) <=  \
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    27
\     (UN i: lessThan n. (A i) Int (- A (Suc i)))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    29
by (Asm_simp_tac 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    30
by (simp_tac (simpset() addsimps [lessThan_Suc]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
qed "UN_Int_Compl_subset";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
(*Converse inclusion fails.*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    36
Goal "(INT i: lessThan n. -A i Un A (Suc i))  <= \
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    37
\     (INT i: lessThan n. -A i) Un A n";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    39
by (Asm_simp_tac 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    40
by (asm_simp_tac (simpset() addsimps [lessThan_Suc]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
qed "INT_Un_Compl_subset";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    45
(*Specialized rewriting*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    46
Goal "A 0 Int (-(A n) Int (INT i: lessThan n. -A i Un A (Suc i))) = {}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
by (blast_tac (claset() addIs [gr0I]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    48
		        addDs [impOfSubs INT_Un_Compl_subset]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    49
val lemma = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
(*Reverse direction makes it harder to invoke the ind hyp*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    52
Goal "(INT i: atMost n. A i) = \
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    53
\         A 0 Int (INT i: lessThan n. -A i Un A(Suc i))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
by (asm_simp_tac
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    57
    (simpset() addsimps Int_ac @ [Int_Un_distrib, Int_Un_distrib2, lemma,
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    58
				  lessThan_Suc, atMost_Suc]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
qed "INT_le_equals_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    61
Goal "(INT i: atMost (Suc n). A i) = \
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    62
\     A 0 Int (INT i: atMost n. -A i Un A(Suc i))";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    63
by (simp_tac (simpset() addsimps [lessThan_Suc_atMost, INT_le_equals_Int]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
qed "INT_le_Suc_equals_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
(*The final deadlock example*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    68
val [zeroprem, allprem] = goalw thy [stable_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 5648
diff changeset
    69
    "[| F : (A 0 Int A (Suc n)) co (A 0);  \
281d44905cab made many specification operators infix
paulson
parents: 5648
diff changeset
    70
\       ALL i: atMost n. F : (A(Suc i) Int A i) co \
281d44905cab made many specification operators infix
paulson
parents: 5648
diff changeset
    71
\                            (-A i Un A(Suc i)) |] \
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5536
diff changeset
    72
\    ==> F : stable (INT i: atMost (Suc n). A i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    73
by (rtac ([zeroprem, allprem RS ball_constrains_INT] MRS 
7127
48e235179ffb added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents: 6536
diff changeset
    74
	  constrains_Int RS constrains_weaken) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    75
by (simp_tac (simpset() addsimps [Collect_le_Int_equals, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    76
				  Int_assoc, INT_absorb]) 1);
5536
130f3d891fb2 tidying and deleting needless parentheses
paulson
parents: 5502
diff changeset
    77
by (simp_tac (simpset() addsimps [INT_le_Suc_equals_Int]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    78
result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    79