| author | nipkow | 
| Tue, 25 Mar 2025 09:40:45 +0100 | |
| changeset 82340 | dc8c25634697 | 
| parent 75916 | b6589c8ccadd | 
| permissions | -rw-r--r-- | 
| 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>UNITY--Chandy and Misra's UNITY formalism\<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 | The book \<^emph>\<open>Parallel Program Design: A Foundation\<close> by Chandy and Misra | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 8 | (Addison-Wesley, 1988) presents the UNITY formalism. UNITY consists of an | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 9 | abstract programming language of guarded assignments and a calculus for | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 10 | reasoning about such programs. Misra's 1994 paper "A Logic for Concurrent | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 11 | Programming" presents New UNITY, giving more elegant foundations for a more | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 12 | general class of languages. In recent work, Chandy and Sanders have proposed | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 13 | new methods for reasoning about systems composed of many components. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 14 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 15 | This directory formalizes these new ideas for UNITY. The Isabelle examples | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 16 | may seem strange to UNITY traditionalists. Hand UNITY proofs tend to be | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 17 | written in the forwards direction, as in informal mathematics, while | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 18 | Isabelle works best in a backwards (goal-directed) style. Programs are | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 19 | expressed as sets of commands, where each command is a relation on states. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 20 | Quantification over commands using \<^verbatim>\<open>[]\<close> is easily expressed. At present, | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 21 | there are no examples of quantification using \<^verbatim>\<open>||\<close>. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 22 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 23 | A UNITY assertion denotes the set of programs satisfying it, as in the | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 24 | propositions-as-types paradigm. The resulting style is readable if | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 25 | unconventional. | 
| 
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 | Safety proofs (invariants) are often proved automatically. Progress proofs | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 28 | involving ENSURES can sometimes be proved automatically. The level of | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 29 | automation appears to be about the same as in HOL-UNITY by Flemming Andersen | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 30 | et al. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 31 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 32 | The directory \<^dir>\<open>Simple\<close> presents a few examples, mostly taken from Misra's | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 33 | 1994 paper, involving single programs. The directory \<^dir>\<open>Comp\<close> presents | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 34 | examples of proofs involving program composition. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 35 | \<close> | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 36 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 37 | end |