| author | wenzelm | 
| Fri, 03 Nov 2023 10:30:51 +0100 | |
| changeset 78888 | 95bbf9a576b3 | 
| 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>Auth--The Inductive Approach to Verifying Security Protocols\<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  | 
Cryptographic protocols are of major importance, especially with the growing  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
use of the Internet. This directory demonstrates the ``inductive method'' of  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
protocol verification, which is described in papers:  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
\<^url>\<open>http://www.cl.cam.ac.uk/users/lcp/papers/protocols.html\<close>. The operational  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
semantics of protocol participants is defined inductively.  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
This directory contains proofs concerning:  | 
| 
 
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  | 
\<^item> three versions of the Otway-Rees protocol  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
\<^item> the Needham-Schroeder shared-key protocol  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
\<^item> the Needham-Schroeder public-key protocol (original and with Lowe's  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
modification)  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
\<^item> two versions of Kerberos: the simplified form published in the BAN paper  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
and also the full protocol (Kerberos IV)  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
\<^item> three versions of the Yahalom protocol, including a bad one that  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
illustrates the purpose of the Oops rule  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
\<^item> a novel recursive authentication protocol  | 
| 
 
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  | 
\<^item> the Internet protocol TLS  | 
| 
 
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  | 
\<^item> The certified e-mail protocol of Abadi et al.  | 
| 
 
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  | 
Frederic Blanqui has contributed a theory of guardedness, which is  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
demonstrated by proofs of some roving agent protocols.  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
\<close>  | 
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
b6589c8ccadd
discontinued special support for README.html (which was hardly ever used in the past 2 decades);
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
end  |