Import of AFP for Isabelle 2021.

This commit is contained in:
Achim D. Brucker 2021-03-01 05:47:16 +00:00
parent 0ad8d1fed7
commit ef59bf6a36
7 changed files with 149 additions and 145 deletions

View File

@ -41,6 +41,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
section \<open>Parallel Compositionality of Security Protocols\<close>
text \<open>\label{sec:Parallel-Compositionality}\<close>
theory Parallel_Compositionality
imports Typing_Result Labeled_Strands
begin

View File

@ -38,6 +38,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
section \<open>Stateful Protocol Compositionality\<close>
text \<open>\label{Stateful-Compositionality}\<close>
theory Stateful_Compositionality
imports Stateful_Typing Parallel_Compositionality Labeled_Stateful_Strands

View File

@ -37,6 +37,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
section \<open>Extending the Typing Result to Stateful Constraints\<close>
text \<open>\label{sec:Stateful-Typing}\<close>
theory Stateful_Typing
imports Typing_Result Stateful_Strands

View File

@ -37,6 +37,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
section \<open>The Typing Result\<close>
text \<open>\label{sec:Typing-Result}\<close>
theory Typing_Result
imports Typed_Model
@ -344,7 +345,7 @@ proof
unfolding \<I>_def wf\<^sub>t\<^sub>r\<^sub>m_def by simp
} hence props: "\<I> v = t \<Longrightarrow> \<Gamma> (Var v) = \<Gamma> t \<and> public_ground_wf_term t" for v t by metis
have "\<I> v \<noteq> Var v" for v using props pgwt_ground by force
have "\<I> v \<noteq> Var v" for v using props pgwt_ground by (simp add: empty_fv_not_var)
hence "subst_domain \<I> = UNIV" by auto
moreover have "ground (subst_range \<I>)" by (simp add: props pgwt_ground)
ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<I>" by metis
@ -2099,7 +2100,7 @@ proof -
by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append)
hence "Ana (Fun f T \<cdot> \<I>) = (K \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<I>, M \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<I>)"
using Ana_t snoc.prems(1)
unfolding Ana_invar_subst_def by force
unfolding Ana_invar_subst_def by blast
ultimately show ?thesis using Decomp t by (auto simp add: decomp_ik)
qed
thus ?case using IH unfolding subst_apply_extstrand_def by simp
@ -2130,7 +2131,7 @@ proof -
using t Decomp snoc.prems(2)
by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append)
hence "Ana (Fun f T \<cdot> \<I>) = (K \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<I>, M \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<I>)"
using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by force
using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast
ultimately show ?thesis using Decomp t by (auto simp add: decomp_assignment_rhs_empty)
qed
thus ?case using IH unfolding subst_apply_extstrand_def by simp

View File

@ -1,6 +1,6 @@
\documentclass[10pt,DIV16,a4paper,abstract=true,twoside=semi,openright]
{scrreprt}
\usepackage[USenglish]{babel}
\usepackage[english]{babel}
\usepackage[numbers, sort&compress]{natbib}
\usepackage{isabelle,isabellesym}
\usepackage{booktabs}
@ -22,8 +22,6 @@
\begingroup%
\def\isacharunderscore{\textunderscore}%
\section{#1 (\thy)}%
\def\isacharunderscore{-}%
\expandafter\label{sec:\isabellecontext}%
\endgroup%
}
@ -69,9 +67,9 @@
The rest of this document is automatically generated from the formalization in Isabelle/HOL, i.e., all content is checked by Isabelle.
The formalization presented in this entry is described in more detail in several publications:
\begin{itemize}
\item The typing result (\autoref{sec:Typing{-}Result} ``Typing\_Result'') for stateless protocols, the TLS formalization (\autoref{sec:Example{-}TLS} ``Example\_TLS''), and the theories depending on those (see \autoref{fig:session-graph}) are described in~\cite{hess.ea:formalizing:2017} and~\cite[chapter 3]{hess:typing:2018}.
\item The typing result for stateful protocols (\autoref{sec:Stateful{-}Typing} ``Stateful\_Typing'') and the keyserver example (\autoref{sec:Example{-}Keyserver} ``Example\_Keyserver'') are described in~\cite{hess.ea:typing:2018} and~\cite[chapter 4]{hess:typing:2018}.
\item The results on parallel composition for stateless protocols (\autoref{sec:Parallel{-}Compositionality} ``Parallel\_Compositionality'') and stateful protocols (\autoref{sec:Stateful{-}Compositionality} ``Stateful\_Compositionality'') are described in~\cite{hess.ea:stateful:2018} and~\cite[chapter 5]{hess:typing:2018}.
\item The typing result (\autoref{sec:Typing-Result} ``Typing\_Result'') for stateless protocols, the TLS formalization (\autoref{sec:Example-TLS} ``Example\_TLS''), and the theories depending on those (see \autoref{fig:session-graph}) are described in~\cite{hess.ea:formalizing:2017} and~\cite[chapter 3]{hess:typing:2018}.
\item The typing result for stateful protocols (\autoref{sec:Stateful-Typing} ``Stateful\_Typing'') and the keyserver example (\autoref{sec:Example-Keyserver} ``Example\_Keyserver'') are described in~\cite{hess.ea:typing:2018} and~\cite[chapter 4]{hess:typing:2018}.
\item The results on parallel composition for stateless protocols (\autoref{sec:Parallel-Compositionality} ``Parallel\_Compositionality'') and stateful protocols (\autoref{sec:Stateful-Compositionality} ``Stateful\_Compositionality'') are described in~\cite{hess.ea:stateful:2018} and~\cite[chapter 5]{hess:typing:2018}.
\end{itemize}
Overall, the structure of this document follows the theory dependencies (see \autoref{fig:session-graph}): we start with introducing the technical preliminaries of our formalization (\autoref{cha:preliminaries}).
Next, we introduce the typing results in \autoref{cha:typing} and \autoref{cha:stateful-typing}.
@ -134,8 +132,8 @@ This work is an extension of the work described in~\cite{hess.ea:stateful:2018}
\chapter{Examples}
\label{cha:examples}
In this chapter, we present two examples illustrating our results:
In \autoref{sec:Example{-}TLS} we show that the TLS example from~\cite{hess.ea:formalizing:2017} is type-flaw resistant.
In \autoref{sec:Example{-}Keyserver} we show that the keyserver examples from~\cite{hess.ea:typing:2018,hess.ea:stateful:2018} are also type-flaw resistant and that the steps of the composed keyserver protocol from~\cite{hess.ea:stateful:2018} satisfy our conditions for protocol composition.
In \autoref{sec:Example-TLS} we show that the TLS example from~\cite{hess.ea:formalizing:2017} is type-flaw resistant.
In \autoref{sec:Example-Keyserver} we show that the keyserver examples from~\cite{hess.ea:typing:2018,hess.ea:stateful:2018} are also type-flaw resistant and that the steps of the composed keyserver protocol from~\cite{hess.ea:stateful:2018} satisfy our conditions for protocol composition.
\input{Example_TLS.tex}
\input{Example_Keyserver.tex}

View File

@ -38,6 +38,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
section \<open>The Keyserver Example\<close>
text \<open>\label{sec:Example-Keyserver}\<close>
theory Example_Keyserver
imports "../Stateful_Compositionality"
begin

View File

@ -37,6 +37,7 @@ OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
section \<open>Proving Type-Flaw Resistance of the TLS Handshake Protocol\<close>
text \<open>\label{sec:Example-TLS}\<close>
theory Example_TLS
imports "../Typed_Model"
begin