src/HOL/UNITY/Simple/Common.thy
author paulson
Thu, 30 Jan 2003 18:08:09 +0100
changeset 13797 baefae13ad37
parent 13785 e2fcd88be55d
child 13806 fd40c9d9076b
permissions -rw-r--r--
conversion of UNITY theories to new-style
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    13
theory Common = UNITY_Main:
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
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    20
  fmono: "m <= n ==> ftime m <= ftime n"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    21
  gmono: "m <= n ==> gtime m <= gtime n"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    22
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    23
  fasc:  "m <= ftime n"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    24
  gasc:  "m <= 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"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    31
    "maxfg m == {t. t <= max (ftime m) (gtime m)}"
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:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    36
     "[| ALL m. F : {m} Co (maxfg m); n: common |]  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    37
      ==> F : Stable (atMost n)"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    38
apply (drule_tac M = "{t. t<=n}" in Elimination_sing)
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)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    41
apply (blast intro: order_eq_refl fmono gmono le_trans)
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:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    45
     "[| Init F <= atMost n;   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    46
         ALL m. F : {m} Co (maxfg m); n: common |]  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    47
      ==> F : Always (atMost n)"
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
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    53
lemma "SKIP : {m} co (maxfg m)"
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*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    57
lemma "mk_program (UNIV, {range(%t.(t,ftime t)), range(%t.(t,gtime t))}, UNIV)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    58
       : {m} co (maxfg m)"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    59
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
    60
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    61
(*This one is  t := max (ftime t) (gtime t)*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    62
lemma  "mk_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}, UNIV)  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    63
       : {m} co (maxfg m)"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    64
by (simp add: constrains_def maxfg_def max_def gasc)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    65
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    66
(*This one is  t := t+1 if t <max (ftime t) (gtime t) *)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    67
lemma  "mk_program   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    68
          (UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }, UNIV)   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    69
       : {m} co (maxfg m)"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    70
by (simp add: constrains_def maxfg_def max_def gasc)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    71
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    72
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    73
(*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
    74
  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
    75
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
(*** Progress under weak fairness ***)
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
declare atMost_Int_atLeast [simp]
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    80
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    81
lemma leadsTo_common_lemma:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    82
     "[| ALL m. F : {m} Co (maxfg m);  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    83
         ALL m: lessThan n. F : {m} LeadsTo (greaterThan m);  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    84
         n: common |]   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    85
      ==> F : (atMost n) LeadsTo common"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    86
apply (rule LeadsTo_weaken_R)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    87
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
    88
apply (simp_all (no_asm_simp))
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    89
apply (rule_tac [2] subset_refl)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    90
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
    91
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    92
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    93
(*The "ALL m: -common" form echoes CMT6.*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    94
lemma leadsTo_common:
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    95
     "[| ALL m. F : {m} Co (maxfg m);  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    96
         ALL m: -common. F : {m} LeadsTo (greaterThan m);  
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    97
         n: common |]   
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    98
      ==> F : (atMost (LEAST n. n: common)) LeadsTo common"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
    99
apply (rule leadsTo_common_lemma)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   100
apply (simp_all (no_asm_simp))
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   101
apply (erule_tac [2] LeastI)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   102
apply (blast dest!: not_less_Least)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   103
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 11195
diff changeset
   104
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
   105
end