src/HOL/UNITY/Deadlock.ML
author paulson
Thu, 02 Jul 1998 16:53:55 +0200
changeset 5111 8f4b72f0c15d
parent 5069 3ea049f7979d
child 5232 e5a7cdd07ea5
permissions -rw-r--r--
Uncurried functions LeadsTo and reach Deleted leading parameters thanks to new Goal command
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     1
(*  Title:      HOL/UNITY/Traces
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]
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    12
    "[| constrains Acts (A Int B) A;  constrains Acts (B Int A) B |] \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
\           ==> stable Acts (A Int B)";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    18
Goal "{i. i < Suc n} = insert n {i. i < n}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
by (blast_tac (claset() addSEs [less_SucE] addIs [less_SucI]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
qed "Collect_less_Suc_insert";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    23
Goal "{i. i <= Suc n} = insert (Suc n) {i. i <= n}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
by (blast_tac (claset() addSEs [le_SucE] addIs [le_SucI]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
qed "Collect_le_Suc_insert";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
(*a simplification step*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    29
Goal "(INT i:{i. i <= n}. A(Suc i) Int A i) = \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    30
\         (INT i:{i. i <= Suc n}. A i)";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Collect_le_Suc_insert])));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
by (blast_tac (claset() addEs [le_SucE] addSEs [equalityE]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
qed "Collect_le_Int_equals";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
(*Dual of the required property.  Converse inclusion fails.*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    38
Goal "(UN i:{i. i < n}. A i) Int Compl (A n) <=  \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    39
\         (UN i:{i. i < n}. (A i) Int Compl (A(Suc i)))";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    40
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
by (simp_tac (simpset() addsimps [Collect_less_Suc_insert]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
qed "UN_Int_Compl_subset";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    45
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    46
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
(*Converse inclusion fails.*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    48
Goal "(INT i:{i. i < n}. Compl(A i) Un A (Suc i))  <= \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    49
\         (INT i:{i. i < n}. Compl(A i)) Un A n";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    52
by (asm_simp_tac (simpset() addsimps [Collect_less_Suc_insert]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    53
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
qed "INT_Un_Compl_subset";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
(*Specialized rewriting*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    58
Goal "A 0 Int (Compl (A n) Int \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
\                  (INT i:{i. i < n}. Compl(A i) Un A (Suc i))) = {}";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
by (blast_tac (claset() addIs [gr0I]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    61
		        addDs [impOfSubs INT_Un_Compl_subset]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    62
val lemma = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    63
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
(*Reverse direction makes it harder to invoke the ind hyp*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    65
Goal "(INT i:{i. i <= n}. A i) = \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66
\         A 0 Int (INT i:{i. i < n}. Compl(A i) Un A(Suc i))";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
by (induct_tac "n" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    68
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    69
by (asm_simp_tac
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    70
    (simpset() addsimps (Int_ac @
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    71
			 [Int_Un_distrib, Int_Un_distrib2, lemma,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    72
			  Collect_less_Suc_insert, Collect_le_Suc_insert])) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    73
qed "INT_le_equals_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    74
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    75
Goal "(INT i:{i. i <= Suc n}. A i) = \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    76
\         A 0 Int (INT i:{i. i <= n}. Compl(A i) Un A(Suc i))";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    77
by (simp_tac (simpset() addsimps [le_eq_less_Suc RS sym, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    78
				  INT_le_equals_Int]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    79
qed "INT_le_Suc_equals_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    80
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    81
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    82
(*The final deadlock example*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    83
val [zeroprem, allprem] = goalw thy [stable_def]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    84
    "[| constrains Acts (A 0 Int A (Suc n)) (A 0);  \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    85
\       ALL i:{i. i <= n}. constrains Acts (A(Suc i) Int A i) \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    86
\                                         (Compl(A i) Un A(Suc i)) |] \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    87
\    ==> stable Acts (INT i:{i. i <= Suc n}. A i)";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    88
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    89
by (rtac ([zeroprem, allprem RS ball_constrains_INT] MRS 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    90
    constrains_Int RS constrains_weaken) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    91
by (simp_tac (simpset() addsimps [Collect_le_Int_equals, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    92
				  Int_assoc, INT_absorb]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    93
by (simp_tac (simpset() addsimps ([INT_le_Suc_equals_Int])) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    94
result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    95