src/HOL/TLA/Stfun.thy
author wenzelm
Tue, 24 Jul 2007 19:44:32 +0200
changeset 23961 9e7e1e309ebd
parent 21624 6f79647cf536
child 35108 e384e27c229f
permissions -rw-r--r--
Multithreading in Poly/ML (version 5.1).
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     1
(*
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     2
    File:        TLA/Stfun.thy
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     3
    ID:          $Id$
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
    Author:      Stephan Merz
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
     5
    Copyright:   1998 University of Munich
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
     6
*)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     7
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
     8
header {* States and state functions for TLA as an "intensional" logic *}
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     9
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    10
theory Stfun
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    11
imports Intensional
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    12
begin
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    13
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    14
typedecl state
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    15
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    16
instance state :: world ..
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    17
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    18
types
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    19
  'a stfun = "state => 'a"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    20
  stpred  = "bool stfun"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    21
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    22
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    23
consts
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    24
  (* Formalizing type "state" would require formulas to be tagged with
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    25
     their underlying state space and would result in a system that is
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    26
     much harder to use. (Unlike Hoare logic or Unity, TLA has quantification
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    27
     over state variables, and therefore one usually works with different
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    28
     state spaces within a single specification.) Instead, "state" is just
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    29
     an anonymous type whose only purpose is to provide "Skolem" constants.
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    30
     Moreover, we do not define a type of state variables separate from that
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    31
     of arbitrary state functions, again in order to simplify the definition
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    32
     of flexible quantification later on. Nevertheless, we need to distinguish
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    33
     state variables, mainly to define the enabledness of actions. The user
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    34
     identifies (tuples of) "base" state variables in a specification via the
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    35
     "meta predicate" basevars, which is defined here.
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    36
  *)
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    37
  stvars    :: "'a stfun => bool"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    38
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    39
syntax
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    40
  "PRED"    :: "lift => 'a"                          ("PRED _")
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    41
  "_stvars" :: "lift => bool"                        ("basevars _")
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    42
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    43
translations
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    44
  "PRED P"   =>  "(P::state => _)"
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    45
  "_stvars"  ==  "stvars"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    46
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    47
defs
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    48
  (* Base variables may be assigned arbitrary (type-correct) values.
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    49
     Note that vs may be a tuple of variables. The correct identification
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    50
     of base variables is up to the user who must take care not to
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    51
     introduce an inconsistency. For example, "basevars (x,x)" would
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    52
     definitely be inconsistent.
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    53
  *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    54
  basevars_def:  "stvars vs == range vs = UNIV"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    55
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    56
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    57
lemma basevars: "!!vs. basevars vs ==> EX u. vs u = c"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    58
  apply (unfold basevars_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    59
  apply (rule_tac b = c and f = vs in rangeE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    60
   apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    61
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    62
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    63
lemma base_pair1: "!!x y. basevars (x,y) ==> basevars x"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    64
  apply (simp (no_asm) add: basevars_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    65
  apply (rule equalityI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    66
   apply (rule subset_UNIV)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    67
  apply (rule subsetI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    68
  apply (drule_tac c = "(xa, arbitrary) " in basevars)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    69
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    70
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    71
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    72
lemma base_pair2: "!!x y. basevars (x,y) ==> basevars y"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    73
  apply (simp (no_asm) add: basevars_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    74
  apply (rule equalityI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    75
   apply (rule subset_UNIV)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    76
  apply (rule subsetI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    77
  apply (drule_tac c = "(arbitrary, xa) " in basevars)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    78
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    79
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    80
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    81
lemma base_pair: "!!x y. basevars (x,y) ==> basevars x & basevars y"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    82
  apply (rule conjI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    83
  apply (erule base_pair1)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    84
  apply (erule base_pair2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    85
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    86
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    87
(* Since the unit type has just one value, any state function can be
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    88
   regarded as "base". The following axiom can sometimes be useful
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    89
   because it gives a trivial solution for "basevars" premises.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    90
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    91
lemma unit_base: "basevars (v::unit stfun)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    92
  apply (unfold basevars_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    93
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    94
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    95
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    96
lemma baseE: "[| basevars v; !!x. v x = c ==> Q |] ==> Q"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    97
  apply (erule basevars [THEN exE])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    98
  apply blast
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    99
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   100
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   101
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   102
(* -------------------------------------------------------------------------------
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   103
   The following shows that there should not be duplicates in a "stvars" tuple:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   104
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   105
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   106
lemma "!!v. basevars (v::bool stfun, v) ==> False"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   107
  apply (erule baseE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   108
  apply (subgoal_tac "(LIFT (v,v)) x = (True, False)")
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   109
   prefer 2
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   110
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   111
  apply simp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   112
  done
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
   113
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
   114
end