src/HOL/TLA/README.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 75916 b6589c8ccadd
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
75916
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     1
theory README imports Main
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     2
begin
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     3
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     4
section \<open>TLA: Lamport's Temporal Logic of Actions\<close>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     5
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     6
text \<open>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     7
  TLA \<^url>\<open>http://www.research.digital.com/SRC/personal/Leslie_Lamport/tla/tla.html\<close>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     8
  is a linear-time temporal logic introduced by Leslie Lamport in \<^emph>\<open>The
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
     9
  Temporal Logic of Actions\<close> (ACM TOPLAS 16(3), 1994, 872-923). Unlike other
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    10
  temporal logics, both systems and properties are represented as logical
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    11
  formulas, and logical connectives such as implication, conjunction, and
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    12
  existential quantification represent structural relations such as
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    13
  refinement, parallel composition, and hiding. TLA has been applied to
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    14
  numerous case studies.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    15
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    16
  This directory formalizes TLA in Isabelle/HOL, as follows:
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    17
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    18
    \<^item> \<^file>\<open>Intensional.thy\<close> prepares the ground by introducing basic syntax for
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    19
      "lifted", possible-world based logics.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    20
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    21
    \<^item> \<^file>\<open>Stfun.thy\<close> and \<^file>\<open>Action.thy\<close> represent the state and transition
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    22
      level formulas of TLA, evaluated over single states and pairs of states.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    23
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    24
    \<^item> \<^file>\<open>Init.thy\<close> introduces temporal logic and defines conversion functions
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    25
      from nontemporal to temporal formulas.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    26
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    27
    \<^item> \<^file>\<open>TLA.thy\<close> axiomatizes proper temporal logic.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    28
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    29
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    30
  Please consult the \<^emph>\<open>design notes\<close>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    31
  \<^url>\<open>http://www.pst.informatik.uni-muenchen.de/~merz/isabelle/IsaTLADesign.ps\<close>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    32
  for further information regarding the setup and use of this encoding of TLA.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    33
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    34
  The theories are accompanied by a small number of examples:
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    35
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    36
    \<^item> \<^dir>\<open>Inc\<close>: Lamport's \<^emph>\<open>increment\<close> example, a standard TLA benchmark,
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    37
      illustrates an elementary TLA proof.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    38
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    39
    \<^item> \<^dir>\<open>Buffer\<close>: a proof that two buffers in a row implement a single buffer,
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    40
      uses a simple refinement mapping.
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    41
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    42
    \<^item> \<^dir>\<open>Memory\<close>: a verification of (the untimed part of) Broy and Lamport's
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    43
    \<^emph>\<open>RPC-Memory\<close> case study, more fully explained in LNCS 1169 (the \<^emph>\<open>TLA
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    44
    solution\<close> is available separately from
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    45
    \<^url>\<open>http://www.pst.informatik.uni-muenchen.de/~merz/papers/RPCMemory.html\<close>).
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    46
\<close>
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    47
b6589c8ccadd discontinued special support for README.html (which was hardly ever used in the past 2 decades);
wenzelm
parents:
diff changeset
    48
end