src/HOL/UNITY/Simple/Deadlock.thy
author paulson
Mon, 26 Jul 2004 17:34:52 +0200
changeset 15077 89840837108e
parent 13806 fd40c9d9076b
child 16417 9bc16273c2d4
permissions -rw-r--r--
converting Hyperreal/Transcendental to Isar script
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     1
(*  Title:      HOL/UNITY/Deadlock
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     2
    ID:         $Id$
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     4
    Copyright   1998  University of Cambridge
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     5
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     6
Deadlock examples from section 5.6 of 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     7
    Misra, "A Logic for Concurrent Programming", 1994
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     8
*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
     9
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    10
theory Deadlock = UNITY:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    11
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    12
(*Trivial, two-process case*)
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    13
lemma "[| F \<in> (A \<inter> B) co A;  F \<in> (B \<inter> A) co B |] ==> F \<in> stable (A \<inter> B)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    14
by (unfold constrains_def stable_def, blast)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    15
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    16
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    17
(*a simplification step*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    18
lemma Collect_le_Int_equals:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    19
     "(\<Inter>i \<in> atMost n. A(Suc i) \<inter> A i) = (\<Inter>i \<in> atMost (Suc n). A i)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    20
apply (induct_tac "n")
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    21
apply (auto simp add: atMost_Suc)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    22
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    23
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    24
(*Dual of the required property.  Converse inclusion fails.*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    25
lemma UN_Int_Compl_subset:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    26
     "(\<Union>i \<in> lessThan n. A i) \<inter> (- A n) \<subseteq>   
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    27
      (\<Union>i \<in> lessThan n. (A i) \<inter> (- A (Suc i)))"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    28
apply (induct_tac "n", simp)
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    29
apply (simp add: lessThan_Suc, blast)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    30
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    31
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    32
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    33
(*Converse inclusion fails.*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    34
lemma INT_Un_Compl_subset:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    35
     "(\<Inter>i \<in> lessThan n. -A i \<union> A (Suc i))  \<subseteq>  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    36
      (\<Inter>i \<in> lessThan n. -A i) \<union> A n"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    37
apply (induct_tac "n", simp)
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    38
apply (simp add: lessThan_Suc, blast)
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    39
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    40
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    41
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    42
(*Specialized rewriting*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    43
lemma INT_le_equals_Int_lemma:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    44
     "A 0 \<inter> (-(A n) \<inter> (\<Inter>i \<in> lessThan n. -A i \<union> A (Suc i))) = {}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    45
by (blast intro: gr0I dest: INT_Un_Compl_subset [THEN subsetD])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    46
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    47
(*Reverse direction makes it harder to invoke the ind hyp*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    48
lemma INT_le_equals_Int:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    49
     "(\<Inter>i \<in> atMost n. A i) =  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    50
      A 0 \<inter> (\<Inter>i \<in> lessThan n. -A i \<union> A(Suc i))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    51
apply (induct_tac "n", simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    52
apply (simp add: Int_ac Int_Un_distrib Int_Un_distrib2
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    53
                 INT_le_equals_Int_lemma lessThan_Suc atMost_Suc)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    54
done
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
lemma INT_le_Suc_equals_Int:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    57
     "(\<Inter>i \<in> atMost (Suc n). A i) =  
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    58
      A 0 \<inter> (\<Inter>i \<in> atMost n. -A i \<union> A(Suc i))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    59
by (simp add: lessThan_Suc_atMost INT_le_equals_Int)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    60
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    61
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    62
(*The final deadlock example*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    63
lemma
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    64
  assumes zeroprem: "F \<in> (A 0 \<inter> A (Suc n)) co (A 0)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    65
      and allprem:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    66
	    "!!i. i \<in> atMost n ==> F \<in> (A(Suc i) \<inter> A i) co (-A i \<union> A(Suc i))"
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    67
  shows "F \<in> stable (\<Inter>i \<in> atMost (Suc n). A i)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    68
apply (unfold stable_def) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    69
apply (rule constrains_Int [THEN constrains_weaken])
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    70
   apply (rule zeroprem) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    71
  apply (rule constrains_INT) 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    72
  apply (erule allprem)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    73
 apply (simp add: Collect_le_Int_equals Int_assoc INT_absorb)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    74
apply (simp add: INT_le_Suc_equals_Int)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    75
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    76
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    77
end