| author | paulson <lp15@cam.ac.uk> | 
| Wed, 01 Feb 2023 12:43:33 +0000 | |
| changeset 77166 | 0fb350e7477b | 
| 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_Guard 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>Protocol-Independent Secrecy Results\<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 | \<^item> date: April 2002 | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 8 | \<^item> author: Frederic Blanqui | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 9 | \<^item> email: blanqui@lri.fr | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 10 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 11 | The current development is built above the HOL (Higher-Order Logic) Isabelle | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 12 | theory and the formalization of protocols introduced by Larry Paulson. More | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 13 | details are in his paper | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 14 | \<^url>\<open>https://www.cl.cam.ac.uk/users/lcp/papers/Auth/jcs.pdf\<close>: \<^emph>\<open>The Inductive | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 15 | approach to verifying cryptographic protocols\<close> (J. Computer Security 6, | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 16 | pages 85-128, 1998). | 
| 
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 | This directory contains a number of files: | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 19 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 20 | \<^item> \<^file>\<open>Extensions.thy\<close> contains extensions of Larry Paulson's files with | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 21 | many useful lemmas. | 
| 
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 | \<^item> \<^file>\<open>Analz.thy\<close> contains an important theorem about the decomposition of | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 24 | analz between pparts (pairs) and kparts (messages that are not pairs). | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 25 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 26 | \<^item> \<^file>\<open>Guard.thy\<close> contains the protocol-independent secrecy theorem for | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 27 | nonces. | 
| 
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 | \<^item> \<^file>\<open>GuardK.thy\<close> is the same for keys. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 30 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 31 | \<^item> \<^file>\<open>Guard_Public.thy\<close> extends \<^file>\<open>Guard.thy\<close> and \<^file>\<open>GuardK.thy\<close> for | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 32 | public-key protocols. | 
| 
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 | \<^item> \<^file>\<open>Guard_Shared.thy\<close> extends \<^file>\<open>Guard.thy\<close> and \<^file>\<open>GuardK.thy\<close> for | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 35 | symmetric-key protocols. | 
| 
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 | \<^item> \<^file>\<open>List_Msg.thy\<close> contains definitions on lists (inside messages). | 
| 
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> \<^file>\<open>P1.thy\<close> contains the definition of the protocol P1 and the proof of | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 40 | its properties (strong forward integrity, insertion resilience, | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 41 | truncation resilience, data confidentiality and non-repudiability). | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 42 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 43 | \<^item> \<^file>\<open>P2.thy\<close> is the same for the protocol P2 | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 44 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 45 | \<^item> \<^file>\<open>Guard_NS_Public.thy\<close> is for Needham-Schroeder-Lowe | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 46 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 47 | \<^item> \<^file>\<open>Guard_OtwayRees.thy\<close> is for Otway-Rees | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 48 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 49 | \<^item> \<^file>\<open>Guard_Yahalom.thy\<close> is for Yahalom | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 50 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 51 | \<^item> \<^file>\<open>Proto.thy\<close> contains a more precise formalization of protocols with | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 52 | rules and a protocol-independent theorem for proving guardness from a | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 53 | preservation property. It also contains the proofs for Needham-Schroeder | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 54 | as an example. | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 55 | \<close> | 
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 56 | |
| 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 wenzelm parents: diff
changeset | 57 | end |