Initial commit, based on AFP entry dated 2020-05-22.

This commit is contained in:
Achim D. Brucker 2020-05-23 15:19:23 +01:00
commit dbbb5ca340
23 changed files with 25157 additions and 0 deletions

38
CITATION Normal file
View File

@ -0,0 +1,38 @@
To cite the use of this formal theory, please use
Andreas V. Hess, Sebastian Mödersheim, and Achim D. Brucker. Stateful
Protocol Composition and Typing. In Archive of Formal Proofs, 2020.
http://www.isa-afp.org/entries/tateful_Protocol_Composition_and_Typing.html,
Formal proof development
A BibTeX entry for LaTeX users is
@Article{ hess.ea:stateful:2020,
abstract= {We provide in this AFP entry several relative soundness
results for security protocols. In particular, we prove
typing and compositionality results for stateful protocols
(i.e., protocols with mutable state that may span several
sessions), and that focuses on reachability properties. Such
results are useful to simplify protocol verification by
reducing it to a simpler problem: Typing results give
conditions under which it is safe to verify a protocol in a
typed model where only "well-typed" attacks can occur whereas
compositionality results allow us to verify a composed protocol
by only verifying the component protocols in isolation. The
conditions on the protocols under which the results hold are
furthermore syntactic in nature allowing for full automation.
The foundation presented here is used in another entry to
provide fully automated and formalized security proofs of
stateful protocols.},
author = {Andreas V. Hess and Sebastian M{\"o}dersheim and Achim D. Brucker},
date = {2020-04-08},
file = {https://www.brucker.ch/bibliography/download/2020/hess.ea-stateful-outline-2020.pdf},
filelabel= {Outline},
issn = {2150-914x},
journal = {Archive of Formal Proofs},
month = {apr},
note = {\url{http://www.isa-afp.org/entries/tateful_Protocol_Composition_and_Typing.html}, Formal proof development},
pdf = {https://www.brucker.ch/bibliography/download/2020/hess.ea-stateful-2020.pdf},
title = {Stateful Protocol Composition and Typing},
url = {https://www.brucker.ch/bibliography/abstract/hess.ea-stateful-2020},
year = {2020},
}

30
LICENSE Normal file
View File

@ -0,0 +1,30 @@
Copyright (c) 2015-2020 Technical University Denmark, Denmark
2017-2019 The University of Sheffield, UK
2019-2020 University of Exeter, UK
All rights reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
* Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
* Redistributions in binary form must reproduce the above
copyright notice, this list of conditions and the following
disclaimer in the documentation and/or other materials provided
with the distribution.
* Neither the name of the copyright holders nor the names of its
contributors may be used to endorse or promote products derived
from this software without specific prior written permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.

44
README.md Normal file
View File

@ -0,0 +1,44 @@
# Stateful Protocol Composition and Typing (
This git repository contains a local mirror of
[Stateful Protocol Composition and Typing](https://www.isa-afp.org/entries/Stateful_Protocol_Composition_and_Typing.html)
entry of the
[Archive of Formal Proofs (AFP)](https://www.isa-afp.org).
The official AFP releases are tagged. Additionally, this repository
may contain extensions (i.e., a development version) that may be
submitted (as an update of the Stateful Protocol Composition and Typing entry) at a later stage.
## Installation
```console
achim@logicalhacking:~$ isabelle build -D Stateful_Protocol_Composition_and_Typing.html
```
## Authors
* Andreas V. Hess
* [Sebastian Mödersheim](https://people.compute.dtu.dk/samo/)
* [Achim D. Brucker](http://www.brucker.ch/)
## License
This project is licensed under a 3-clause BSD-style license.
SPDX-License-Identifier: BSD-3-Clause
## Master Repository
The master git repository for this project is hosted by the [Software
Assurance & Security Research Team](https://logicalhacking.com) at
<https://git.logicalhacking.com/afp-mirror/Stateful_Protocol_Composition_and_Typing.html>.
## Publications
* Andreas V. Hess, Sebastian Mödersheim, and Achim D. Brucker. Stateful
Protocol Composition and Typing. In Archive of Formal Proofs, 2020.
http://www.isa-afp.org/entries/tateful_Protocol_Composition_and_Typing.html,
Formal proof development
* Andreas V. Hess, Sebastian A. Mödersheim, and Achim D. Brucker. Stateful
Protocol Composition. In ESORICS. Lecture Notes in Computer Science (11098),
pages 427-446, Springer-Verlag, 2018. doi:[10.1007/978-3-319-99073-6](https://dx.doi.org/10.1007/978-3-319-99073-6)

View File

@ -0,0 +1,5 @@
theory Examples
imports "examples/Example_Keyserver"
"examples/Example_TLS"
begin
end

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,906 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2018-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Labeled_Stateful_Strands.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>Labeled Stateful Strands\<close>
theory Labeled_Stateful_Strands
imports Stateful_Strands Labeled_Strands
begin
subsection \<open>Definitions\<close>
text\<open>Syntax for stateful strand labels\<close>
abbreviation Star_step ("\<langle>\<star>, _\<rangle>") where
"\<langle>\<star>, (s::('a,'b) stateful_strand_step)\<rangle> \<equiv> (\<star>, s)"
abbreviation LabelN_step ("\<langle>_, _\<rangle>") where
"\<langle>(l::'a), (s::('b,'c) stateful_strand_step)\<rangle> \<equiv> (ln l, s)"
text\<open>Database projection\<close>
abbreviation dbproj where "dbproj l D \<equiv> filter (\<lambda>d. fst d = l) D"
text\<open>The type of labeled stateful strands\<close>
type_synonym ('a,'b,'c) labeled_stateful_strand_step = "'c strand_label \<times> ('a,'b) stateful_strand_step"
type_synonym ('a,'b,'c) labeled_stateful_strand = "('a,'b,'c) labeled_stateful_strand_step list"
text\<open>Dual strands\<close>
fun dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b,'c) labeled_stateful_strand_step \<Rightarrow> ('a,'b,'c) labeled_stateful_strand_step"
where
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,send\<langle>t\<rangle>) = (l,receive\<langle>t\<rangle>)"
| "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,receive\<langle>t\<rangle>) = (l,send\<langle>t\<rangle>)"
| "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = x"
definition dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t::"('a,'b,'c) labeled_stateful_strand \<Rightarrow> ('a,'b,'c) labeled_stateful_strand"
where
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<equiv> map dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p"
text\<open>Substitution application\<close>
fun subst_apply_labeled_stateful_strand_step::
"('a,'b,'c) labeled_stateful_strand_step \<Rightarrow> ('a,'b) subst \<Rightarrow>
('a,'b,'c) labeled_stateful_strand_step"
(infix "\<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p" 51) where
"(l,s) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta> = (l,s \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>)"
definition subst_apply_labeled_stateful_strand::
"('a,'b,'c) labeled_stateful_strand \<Rightarrow> ('a,'b) subst \<Rightarrow> ('a,'b,'c) labeled_stateful_strand"
(infix "\<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t" 51) where
"S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta> \<equiv> map (\<lambda>x. x \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>) S"
text\<open>Definitions lifted from stateful strands\<close>
abbreviation wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> ik\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation db\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> db\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> db'\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> trms\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation trms_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "trms_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t n S \<equiv> trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n S)"
abbreviation vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> vars\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation vars_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "vars_proj\<^sub>l\<^sub>s\<^sub>s\<^sub>t n S \<equiv> vars\<^sub>s\<^sub>s\<^sub>t (proj_unl n S)"
abbreviation bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> bvars\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
abbreviation fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> fv\<^sub>s\<^sub>s\<^sub>t (unlabel S)"
text\<open>Labeled set-operations\<close>
fun setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p where
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,insert\<langle>t,s\<rangle>) = {(i,t,s)}"
| "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,delete\<langle>t,s\<rangle>) = {(i,t,s)}"
| "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\<langle>_: t \<in> s\<rangle>) = {(i,t,s)}"
| "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\<forall>_\<langle>\<or>\<noteq>: _ \<or>\<notin>: F'\<rangle>) = ((\<lambda>(t,s). (i,t,s)) ` set F')"
| "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = {}"
definition setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t where
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<equiv> \<Union>(setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)"
subsection \<open>Minor Lemmata\<close>
lemma subst_lsst_nil[simp]: "[] \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = []"
by (simp add: subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_cons: "a#A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)#(A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
by (simp add: subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_singleton: "[(l,s)] \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = [(l,s \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)]"
by (simp add: subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_append: "A@B \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)@(B \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
by (simp add: subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_append_inv:
assumes "A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = B1@B2"
shows "\<exists>A1 A2. A = A1@A2 \<and> A1 \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = B1 \<and> A2 \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = B2"
using assms
proof (induction A arbitrary: B1 B2)
case (Cons a A)
note prems = Cons.prems
note IH = Cons.IH
show ?case
proof (cases B1)
case Nil
then obtain b B3 where "B2 = b#B3" "a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> = b" "A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = B3"
using prems subst_lsst_cons by fastforce
thus ?thesis by (simp add: Nil subst_apply_labeled_stateful_strand_def)
next
case (Cons b B3)
hence "a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> = b" "A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> = B3@B2"
using prems by (simp_all add: subst_lsst_cons)
thus ?thesis by (metis Cons_eq_appendI Cons IH subst_lsst_cons)
qed
qed (metis append_is_Nil_conv subst_lsst_nil)
lemma subst_lsst_member[intro]: "x \<in> set A \<Longrightarrow> x \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> \<in> set (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
by (metis image_eqI set_map subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_unlabel_cons: "unlabel ((l,b)#A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = (b \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>)#(unlabel (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
by (simp add: subst_apply_labeled_stateful_strand_def)
lemma subst_lsst_unlabel: "unlabel (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = unlabel A \<cdot>\<^sub>s\<^sub>s\<^sub>t \<delta>"
proof (induction A)
case (Cons a A)
then obtain l b where "a = (l,b)" by (metis surj_pair)
thus ?case
using Cons
by (simp add: subst_apply_labeled_stateful_strand_def subst_apply_stateful_strand_def)
qed simp
lemma subst_lsst_unlabel_member[intro]:
assumes "x \<in> set (unlabel A)"
shows "x \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> \<in> set (unlabel (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>))"
proof -
obtain l where x: "(l,x) \<in> set A" using assms unfolding unlabel_def by moura
thus ?thesis
using subst_lsst_member
by (metis unlabel_def in_set_zipE subst_apply_labeled_stateful_strand_step.simps zip_map_fst_snd)
qed
lemma subst_lsst_prefix:
assumes "prefix B (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
shows "\<exists>C. C \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta> = B \<and> prefix C A"
using assms
proof (induction A rule: List.rev_induct)
case (snoc a A) thus ?case
proof (cases "B = A@[a] \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>")
case False thus ?thesis
using snoc by (auto simp add: subst_lsst_append[of A] subst_lsst_cons)
qed auto
qed simp
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = []"
by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons[simp]:
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,send\<langle>t\<rangle>)#A) = (l,receive\<langle>t\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,receive\<langle>t\<rangle>)#A) = (l,send\<langle>t\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\<langle>a: t \<doteq> s\<rangle>)#A) = (l,\<langle>a: t \<doteq> s\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,insert\<langle>t,s\<rangle>)#A) = (l,insert\<langle>t,s\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,delete\<langle>t,s\<rangle>)#A) = (l,delete\<langle>t,s\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\<langle>a: t \<in> s\<rangle>)#A) = (l,\<langle>a: t \<in> s\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>)#A) = (l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
by (simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append[simp]: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t B"
by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>"
proof -
obtain l x where s: "s = (l,x)" by moura
thus ?thesis by (cases x) (auto simp add: subst_apply_labeled_stateful_strand_def)
qed
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>"
proof (induction S)
case (Cons s S) thus ?case
using Cons dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[of s \<delta>]
by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def)
qed (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \<cdot>\<^sub>s\<^sub>s\<^sub>t \<delta>"
by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<sigma>)#(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>))"
by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst list.simps(9) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_append: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>"
by (metis (no_types) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>) = (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>)@[dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<sigma>]"
by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons list.map(1) map_append
subst_apply_labeled_stateful_strand_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD:
assumes "(l,a) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
shows "\<exists>b. (l,b) \<in> set A \<and> dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)"
using assms
proof (induction A)
case (Cons c A)
hence "(l,a) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \<or> dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p c = (l,a)" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force
thus ?case
proof
assume "(l,a) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" thus ?case using Cons.IH by auto
next
assume a: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p c = (l,a)"
obtain i b where b: "c = (i,b)" by (metis surj_pair)
thus ?case using a by (cases b) auto
qed
qed simp
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv:
assumes "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l, a) = (k, b)"
shows "l = k"
and "a = receive\<langle>t\<rangle> \<Longrightarrow> b = send\<langle>t\<rangle>"
and "a = send\<langle>t\<rangle> \<Longrightarrow> b = receive\<langle>t\<rangle>"
and "(\<nexists>t. a = receive\<langle>t\<rangle> \<or> a = send\<langle>t\<rangle>) \<Longrightarrow> b = a"
proof -
show "l = k" using assms by (cases a) auto
show "a = receive\<langle>t\<rangle> \<Longrightarrow> b = send\<langle>t\<rangle>" using assms by (cases a) auto
show "a = send\<langle>t\<rangle> \<Longrightarrow> b = receive\<langle>t\<rangle>" using assms by (cases a) auto
show "(\<nexists>t. a = receive\<langle>t\<rangle> \<or> a = send\<langle>t\<rangle>) \<Longrightarrow> b = a" using assms by (cases a) auto
qed
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse: "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = A"
proof (induction A)
case (Cons a A)
obtain l b where "a = (l,b)" by (metis surj_pair)
thus ?case using Cons by (cases b) auto
qed simp
lemma vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?case using Cons.IH by (cases b) auto
qed simp
lemma fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?case using Cons.IH by (cases b) auto
qed simp
lemma bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?case using Cons.IH by (cases b) simp+
qed simp
lemma vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \<union> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
by (metis unlabel_Cons(1) vars\<^sub>s\<^sub>s\<^sub>t_Cons)
lemma fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
by (metis unlabel_Cons(1) fv\<^sub>s\<^sub>s\<^sub>t_Cons)
lemma bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \<union> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
by (metis unlabel_Cons(1) bvars\<^sub>s\<^sub>s\<^sub>t_Cons)
lemma bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
by (metis subst_lsst_unlabel bvars\<^sub>s\<^sub>s\<^sub>t_subst)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_member:
assumes "(l,x) \<in> set A"
and "\<not>is_Receive x" "\<not>is_Send x"
shows "(l,x) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
using assms
proof (induction A)
case (Cons a A) thus ?case using assms(2,3) by (cases x) (auto simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
qed simp
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_member:
assumes "x \<in> set (unlabel A)"
and "\<not>is_Receive x" "\<not>is_Send x"
shows "x \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
using assms dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_member[of _ _ A]
by (meson unlabel_in unlabel_mem_has_label)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff:
"(l,send\<langle>t\<rangle>) \<in> set A \<longleftrightarrow> (l,receive\<langle>t\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,receive\<langle>t\<rangle>) \<in> set A \<longleftrightarrow> (l,send\<langle>t\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,\<langle>c: t \<doteq> s\<rangle>) \<in> set A \<longleftrightarrow> (l,\<langle>c: t \<doteq> s\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,insert\<langle>t,s\<rangle>) \<in> set A \<longleftrightarrow> (l,insert\<langle>t,s\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,delete\<langle>t,s\<rangle>) \<in> set A \<longleftrightarrow> (l,delete\<langle>t,s\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,\<langle>c: t \<in> s\<rangle>) \<in> set A \<longleftrightarrow> (l,\<langle>c: t \<in> s\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
"(l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>) \<in> set A \<longleftrightarrow> (l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>) \<in> set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)"
proof (induction A)
case (Cons a A)
obtain j b where a: "a = (j,b)" by (metis surj_pair)
{ case 1 thus ?case by (cases b) (simp_all add: Cons.IH(1) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 2 thus ?case by (cases b) (simp_all add: Cons.IH(2) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 3 thus ?case by (cases b) (simp_all add: Cons.IH(3) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 4 thus ?case by (cases b) (simp_all add: Cons.IH(4) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 5 thus ?case by (cases b) (simp_all add: Cons.IH(5) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 6 thus ?case by (cases b) (simp_all add: Cons.IH(6) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 7 thus ?case by (cases b) (simp_all add: Cons.IH(7) a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) }
qed (simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff:
"send\<langle>t\<rangle> \<in> set (unlabel A) \<longleftrightarrow> receive\<langle>t\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"receive\<langle>t\<rangle> \<in> set (unlabel A) \<longleftrightarrow> send\<langle>t\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"\<langle>c: t \<doteq> s\<rangle> \<in> set (unlabel A) \<longleftrightarrow> \<langle>c: t \<doteq> s\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"insert\<langle>t,s\<rangle> \<in> set (unlabel A) \<longleftrightarrow> insert\<langle>t,s\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"delete\<langle>t,s\<rangle> \<in> set (unlabel A) \<longleftrightarrow> delete\<langle>t,s\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"\<langle>c: t \<in> s\<rangle> \<in> set (unlabel A) \<longleftrightarrow> \<langle>c: t \<in> s\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle> \<in> set (unlabel A) \<longleftrightarrow> \<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1,2)[of _ t A]
dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(3,6)[of _ c t s A]
dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4,5)[of _ t s A]
dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(7)[of _ X F G A]
by (meson unlabel_in unlabel_mem_has_label)+
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all:
"list_all is_Receive (unlabel A) \<Longrightarrow> list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Send (unlabel A) \<Longrightarrow> list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Equality (unlabel A) \<Longrightarrow> list_all is_Equality (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Insert (unlabel A) \<Longrightarrow> list_all is_Insert (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Delete (unlabel A) \<Longrightarrow> list_all is_Delete (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_InSet (unlabel A) \<Longrightarrow> list_all is_InSet (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_NegChecks (unlabel A) \<Longrightarrow> list_all is_NegChecks (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Assignment (unlabel A) \<Longrightarrow> list_all is_Assignment (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Check (unlabel A) \<Longrightarrow> list_all is_Check (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
"list_all is_Update (unlabel A) \<Longrightarrow> list_all is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
proof (induct A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
{ case 1 thus ?case using Cons.hyps(1) a by (cases b) auto }
{ case 2 thus ?case using Cons.hyps(2) a by (cases b) auto }
{ case 3 thus ?case using Cons.hyps(3) a by (cases b) auto }
{ case 4 thus ?case using Cons.hyps(4) a by (cases b) auto }
{ case 5 thus ?case using Cons.hyps(5) a by (cases b) auto }
{ case 6 thus ?case using Cons.hyps(6) a by (cases b) auto }
{ case 7 thus ?case using Cons.hyps(7) a by (cases b) auto }
{ case 8 thus ?case using Cons.hyps(8) a by (cases b) auto }
{ case 9 thus ?case using Cons.hyps(9) a by (cases b) auto }
{ case 10 thus ?case using Cons.hyps(10) a by (cases b) auto }
qed simp_all
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain:
assumes "s \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))"
shows "\<exists>l B s'. (l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,s') \<and> prefix (B@[(l,s')]) A"
using assms
proof (induction A rule: List.rev_induct)
case (snoc a A)
obtain i b where a: "a = (i,b)" by (metis surj_pair)
show ?case using snoc
proof (cases "s \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))")
case False thus ?thesis
using a snoc.prems unlabel_append[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t [a]"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"]
by (cases b) (force simp add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)+
qed auto
qed simp
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst:
assumes "s \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)))"
shows "\<exists>l B s'. (l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s') \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>) \<and> prefix ((B \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)@[(l,s') \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>]) (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
proof -
obtain B l s' where B: "(l,s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,s')" "prefix (B@[(l,s')]) (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain[OF assms] by moura
obtain C where C: "C \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta> = B@[(l,s')]"
using subst_lsst_prefix[OF B(2)] by moura
obtain D u where D: "C = D@[(l,u)]" "D \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta> = B" "[(l,u)] \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta> = [(l, s')]"
using subst_lsst_prefix[OF B(2)] subst_lsst_append_inv[OF C(1)]
by (auto simp add: subst_apply_labeled_stateful_strand_def)
show ?thesis
using B D subst_lsst_cons subst_lsst_singleton
by (metis (no_types, lifting) nth_append_length)
qed
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?case using Cons.IH by (cases b) auto
qed simp
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst_cons:
"trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((l,b)#A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) \<union> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
by (metis subst_lsst_unlabel trms\<^sub>s\<^sub>s\<^sub>t_subst_cons unlabel_Cons(1))
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst:
assumes "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<inter> subst_domain \<theta> = {}"
shows "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta>"
by (metis trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms] subst_lsst_unlabel)
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst':
fixes t::"('a,'b) term" and \<delta>::"('a,'b) subst"
assumes "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
shows "\<exists>s \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S. \<exists>X. set X \<subseteq> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<and> t = s \<cdot> rm_vars (set X) \<delta>"
using assms
proof (induction S)
case (Cons a S)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
hence "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) \<or> t \<in> trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)"
using Cons.prems trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst_cons by fast
thus ?case
proof
assume *: "t \<in> trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)"
show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF *] a by auto
next
assume *: "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
show ?thesis using Cons.IH[OF *] a by auto
qed
qed simp
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'':
fixes t::"('a,'b) term" and \<delta> \<theta>::"('a,'b) subst"
assumes "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta>"
shows "\<exists>s \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S. \<exists>X. set X \<subseteq> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<and> t = s \<cdot> rm_vars (set X) \<delta> \<circ>\<^sub>s \<theta>"
proof -
obtain s where s: "s \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)" "t = s \<cdot> \<theta>" using assms by moura
show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'[OF s(1)] s(2) by auto
qed
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual_subst_cons:
"trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>)) = (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<sigma>)) \<union> (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>)))"
proof -
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?thesis using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_cons[of a A \<sigma>] by (cases b) auto
qed
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_funs_term:
"\<Union>(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S)))) = \<Union>(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t (unlabel S)))"
using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by fast
lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t:
"db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
thus ?case using Cons by (cases b) auto
qed simp
lemma db\<^sub>s\<^sub>s\<^sub>t_unlabel_append:
"db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) I D = db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I D)"
by (metis db\<^sub>s\<^sub>s\<^sub>t_append unlabel_append)
lemma db\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t:
"db'\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>))) \<I> D = db'\<^sub>s\<^sub>s\<^sub>t (unlabel (T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)) \<I> D"
proof (induction T arbitrary: D)
case (Cons x T)
obtain l s where "x = (l,s)" by moura
thus ?case
using Cons
by (cases s) (simp_all add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def)
qed (simp add: unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def)
lemma labeled_list_insert_eq_cases:
"d \<notin> set (unlabel D) \<Longrightarrow> List.insert d (unlabel D) = unlabel (List.insert (i,d) D)"
"(i,d) \<in> set D \<Longrightarrow> List.insert d (unlabel D) = unlabel (List.insert (i,d) D)"
unfolding unlabel_def
by (metis (no_types, hide_lams) List.insert_def image_eqI list.simps(9) set_map snd_conv,
metis in_set_insert set_zip_rightD zip_map_fst_snd)
lemma labeled_list_insert_eq_ex_cases:
"List.insert d (unlabel D) = unlabel (List.insert (i,d) D) \<or>
(\<exists>j. (j,d) \<in> set D \<and> List.insert d (unlabel D) = unlabel (List.insert (j,d) D))"
using labeled_list_insert_eq_cases unfolding unlabel_def
by (metis in_set_impl_in_set_zip2 length_map zip_map_fst_snd)
lemma proj_subst: "proj l (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = proj l A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>"
proof (induction A)
case (Cons a A)
obtain l b where "a = (l,b)" by (metis surj_pair)
thus ?case using Cons unfolding proj_def subst_apply_labeled_stateful_strand_def by force
qed simp
lemma proj_set_subset[simp]:
"set (proj n A) \<subseteq> set A"
unfolding proj_def by auto
lemma proj_proj_set_subset[simp]:
"set (proj n (proj m A)) \<subseteq> set (proj n A)"
"set (proj n (proj m A)) \<subseteq> set (proj m A)"
"set (proj_unl n (proj m A)) \<subseteq> set (proj_unl n A)"
"set (proj_unl n (proj m A)) \<subseteq> set (proj_unl m A)"
unfolding unlabel_def proj_def by auto
lemma proj_in_set_iff:
"(ln i, d) \<in> set (proj i D) \<longleftrightarrow> (ln i, d) \<in> set D"
"(\<star>, d) \<in> set (proj i D) \<longleftrightarrow> (\<star>, d) \<in> set D"
unfolding proj_def by auto
lemma proj_list_insert:
"proj i (List.insert (ln i,d) D) = List.insert (ln i,d) (proj i D)"
"proj i (List.insert (\<star>,d) D) = List.insert (\<star>,d) (proj i D)"
"i \<noteq> j \<Longrightarrow> proj i (List.insert (ln j,d) D) = proj i D"
unfolding List.insert_def proj_def by auto
lemma proj_filter: "proj i [d\<leftarrow>D. d \<notin> set Di] = [d\<leftarrow>proj i D. d \<notin> set Di]"
by (simp_all add: proj_def conj_commute)
lemma proj_list_Cons:
"proj i ((ln i,d)#D) = (ln i,d)#proj i D"
"proj i ((\<star>,d)#D) = (\<star>,d)#proj i D"
"i \<noteq> j \<Longrightarrow> proj i ((ln j,d)#D) = proj i D"
unfolding List.insert_def proj_def by auto
lemma proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t:
"proj l (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A)"
proof (induction A)
case (Cons a A)
obtain k b where "a = (k,b)" by (metis surj_pair)
thus ?case using Cons unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def by (cases b) auto
qed simp
lemma proj_instance_ex:
assumes B: "\<forall>b \<in> set B. \<exists>a \<in> set A. \<exists>\<delta>. b = a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> \<and> P \<delta>"
and b: "b \<in> set (proj l B)"
shows "\<exists>a \<in> set (proj l A). \<exists>\<delta>. b = a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> \<and> P \<delta>"
proof -
obtain a \<delta> where a: "a \<in> set A" "b = a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>" "P \<delta>" using B b proj_set_subset by fast
obtain k b' where b': "b = (k, b')" "k = (ln l) \<or> k = \<star>" using b proj_in_setD by metis
obtain a' where a': "a = (k, a')" using b'(1) a(2) by (cases a) simp_all
show ?thesis using a a' b'(2) unfolding proj_def by auto
qed
lemma proj_dbproj:
"dbproj (ln i) (proj i D) = dbproj (ln i) D"
"dbproj \<star> (proj i D) = dbproj \<star> D"
"i \<noteq> j \<Longrightarrow> dbproj (ln j) (proj i D) = []"
unfolding proj_def by (induct D) auto
lemma dbproj_Cons:
"dbproj i ((i,d)#D) = (i,d)#dbproj i D"
"i \<noteq> j \<Longrightarrow> dbproj j ((i,d)#D) = dbproj j D"
by auto
lemma dbproj_subset[simp]:
"set (unlabel (dbproj i D)) \<subseteq> set (unlabel D)"
unfolding unlabel_def by auto
lemma dbproj_subseq:
assumes "Di \<in> set (subseqs (dbproj k D))"
shows "dbproj k Di = Di" (is ?A)
and "i \<noteq> k \<Longrightarrow> dbproj i Di = []" (is "i \<noteq> k \<Longrightarrow> ?B")
proof -
have *: "set Di \<subseteq> set (dbproj k D)" using subseqs_powset[of "dbproj k D"] assms by auto
thus ?A by (metis filter_True filter_set member_filter subsetCE)
have "\<And>j d. (j,d) \<in> set Di \<Longrightarrow> j = k" using * by auto
moreover have "\<And>j d. (j,d) \<in> set (dbproj i Di) \<Longrightarrow> j = i" by auto
moreover have "\<And>j d. (j,d) \<in> set (dbproj i Di) \<Longrightarrow> (j,d) \<in> set Di" by auto
ultimately show "i \<noteq> k \<Longrightarrow> ?B" by (metis set_empty subrelI subset_empty)
qed
lemma dbproj_subseq_subset:
assumes "Di \<in> set (subseqs (dbproj i D))"
shows "set Di \<subseteq> set D"
by (metis Pow_iff assms filter_set image_eqI member_filter subseqs_powset subsetCE subsetI)
lemma dbproj_subseq_in_subseqs:
assumes "Di \<in> set (subseqs (dbproj i D))"
shows "Di \<in> set (subseqs D)"
using assms in_set_subseqs subseq_filter_left subseq_order.dual_order.trans by blast
lemma proj_subseq:
assumes "Di \<in> set (subseqs (dbproj (ln j) D))" "j \<noteq> i"
shows "[d\<leftarrow>proj i D. d \<notin> set Di] = proj i D"
proof -
have "set Di \<subseteq> set (dbproj (ln j) D)" using subseqs_powset[of "dbproj (ln j) D"] assms by auto
hence "\<And>k d. (k,d) \<in> set Di \<Longrightarrow> k = ln j" by auto
moreover have "\<And>k d. (k,d) \<in> set (proj i D) \<Longrightarrow> k \<noteq> ln j"
using assms(2) unfolding proj_def by auto
ultimately have "\<And>d. d \<in> set (proj i D) \<Longrightarrow> d \<notin> set Di" by auto
thus ?thesis by simp
qed
lemma unlabel_subseqsD:
assumes "A \<in> set (subseqs (unlabel B))"
shows "\<exists>C \<in> set (subseqs B). unlabel C = A"
using assms map_subseqs unfolding unlabel_def by (metis imageE set_map)
lemma unlabel_filter_eq:
assumes "\<forall>(j, p) \<in> set A \<union> B. \<forall>(k, q) \<in> set A \<union> B. p = q \<longrightarrow> j = k" (is "?P (set A)")
shows "[d\<leftarrow>unlabel A. d \<notin> snd ` B] = unlabel [d\<leftarrow>A. d \<notin> B]"
using assms unfolding unlabel_def
proof (induction A)
case (Cons a A)
have "set A \<subseteq> set (a#A)" "{a} \<subseteq> set (a#A)" by auto
hence *: "?P (set A)" "?P {a}" using Cons.prems by fast+
hence IH: "[d\<leftarrow>map snd A . d \<notin> snd ` B] = map snd [d\<leftarrow>A . d \<notin> B]" using Cons.IH by auto
{ assume "snd a \<in> snd ` B"
then obtain b where b: "b \<in> B" "snd a = snd b" by moura
hence "fst a = fst b" using *(2) by auto
hence "a \<in> B" using b by (metis surjective_pairing)
} hence **: "a \<notin> B \<Longrightarrow> snd a \<notin> snd ` B" by metis
show ?case by (cases "a \<in> B") (simp add: ** IH)+
qed simp
lemma subseqs_mem_dbproj:
assumes "Di \<in> set (subseqs D)" "list_all (\<lambda>d. fst d = i) Di"
shows "Di \<in> set (subseqs (dbproj i D))"
using assms
proof (induction D arbitrary: Di)
case (Cons di D)
obtain d j where di: "di = (j,d)" by (metis surj_pair)
show ?case
proof (cases "Di \<in> set (subseqs D)")
case True
hence "Di \<in> set (subseqs (dbproj i D))" using Cons.IH Cons.prems by auto
thus ?thesis using subseqs_Cons by auto
next
case False
then obtain Di' where Di': "Di = di#Di'" using Cons.prems(1)
by (metis (mono_tags, lifting) Un_iff imageE set_append set_map subseqs.simps(2))
hence "Di' \<in> set (subseqs D)" using Cons.prems(1) False
by (metis (no_types, lifting) UnE imageE list.inject set_append set_map subseqs.simps(2))
hence "Di' \<in> set (subseqs (dbproj i D))" using Cons.IH Cons.prems Di' by auto
moreover have "i = j" using Di' di Cons.prems(2) by auto
hence "dbproj i (di#D) = di#dbproj i D" by (simp add: di)
ultimately show ?thesis using Di'
by (metis (no_types, lifting) UnCI image_eqI set_append set_map subseqs.simps(2))
qed
qed simp
lemma unlabel_subst: "unlabel S \<cdot>\<^sub>s\<^sub>s\<^sub>t \<delta> = unlabel (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
unfolding unlabel_def subst_apply_stateful_strand_def subst_apply_labeled_stateful_strand_def
by auto
lemma subterms_subst_lsst:
assumes "\<forall>x \<in> fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\<exists>f. \<sigma> x = Fun f []) \<or> (\<exists>y. \<sigma> x = Var y)"
and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<inter> subst_domain \<sigma> = {}"
shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>)) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<sigma>"
using subterms_subst''[OF assms(1)] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms(2)] unlabel_subst[of S \<sigma>]
by simp
lemma subterms_subst_lsst_ik:
assumes "\<forall>x \<in> fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\<exists>f. \<sigma> x = Fun f []) \<or> (\<exists>y. \<sigma> x = Var y)"
shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<sigma>)) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<sigma>"
using subterms_subst''[OF assms(1)] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel S" \<sigma>] unlabel_subst[of S \<sigma>]
by simp
lemma labeled_stateful_strand_subst_comp:
assumes "range_vars \<delta> \<inter> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = {}"
shows "S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> \<circ>\<^sub>s \<theta> = (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>"
using assms
proof (induction S)
case (Cons s S)
obtain l x where s: "s = (l,x)" by (metis surj_pair)
hence IH: "S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta> \<circ>\<^sub>s \<theta> = (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>" using Cons by auto
have "x \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta> \<circ>\<^sub>s \<theta> = (x \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>"
using s Cons.prems stateful_strand_step_subst_comp[of \<delta> x \<theta>] by auto
thus ?case using s IH by (simp add: subst_apply_labeled_stateful_strand_def)
qed simp
lemma sst_vars_proj_subset[simp]:
"fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)"
"bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A)"
"vars\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> vars\<^sub>s\<^sub>s\<^sub>t (unlabel A)"
using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"]
vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "proj_unl n A"]
unfolding unlabel_def proj_def by auto
lemma trms\<^sub>s\<^sub>s\<^sub>t_proj_subset[simp]:
"trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (unlabel A)" (is ?A)
"trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" (is ?B)
"trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (proj_unl m A)" (is ?C)
proof -
show ?A unfolding unlabel_def proj_def by auto
show ?B using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_proj_set_subset(4)] by metis
show ?C using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_proj_set_subset(3)] by metis
qed
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset:
"trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))" (is ?A)
"trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))" (is ?B)
using trms\<^sub>s\<^sub>s\<^sub>t_mono[of "proj_unl n A" "proj_unl n (A@B)"]
unfolding unlabel_def proj_def by auto
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_suffix_subset:
"trms\<^sub>s\<^sub>s\<^sub>t (unlabel B) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))"
"trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n B) \<subseteq> trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))"
using trms\<^sub>s\<^sub>s\<^sub>t_mono[of "proj_unl n B" "proj_unl n (A@B)"]
unfolding unlabel_def proj_def by auto
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>pD:
assumes p: "p \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a"
shows "fst p = fst a" (is ?P)
and "is_Update (snd a) \<or> is_InSet (snd a) \<or> is_NegChecks (snd a)" (is ?Q)
proof -
obtain l k p' a' where a: "p = (l,p')" "a = (k,a')" by (metis surj_pair)
show ?P using p a by (cases a') auto
show ?Q using p a by (cases a') auto
qed
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil[simp]:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = {}"
by (simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons[simp]:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (x#S) = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \<union> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S"
by (simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma setops\<^sub>s\<^sub>s\<^sub>t_proj_subset:
"setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)"
"setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)"
"setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m (proj n A)) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl m A)"
unfolding unlabel_def proj_def
proof (induction A)
case (Cons a A)
obtain l b where lb: "a = (l,b)" by moura
{ case 1 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 2 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 3 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
qed simp_all
lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset:
"setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))"
"setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))"
unfolding unlabel_def proj_def
proof (induction A)
case (Cons a A)
obtain l b where lb: "a = (l,b)" by moura
{ case 1 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 2 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
qed (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def)
lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_suffix_subset:
"setops\<^sub>s\<^sub>s\<^sub>t (unlabel B) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B))"
"setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n B) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n (A@B))"
unfolding unlabel_def proj_def
proof (induction A)
case (Cons a A)
obtain l b where lb: "a = (l,b)" by moura
{ case 1 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
{ case 2 thus ?case using Cons.IH lb by (cases b) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) }
qed simp_all
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_subset:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj m (proj n A)) \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A)"
unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_prefix_subset:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B)"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n (A@B))"
unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_suffix_subset:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B)"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n (A@B))"
unfolding proj_def setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono:
"set M \<subseteq> set N \<Longrightarrow> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t M \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t N"
by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def)
lemma trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label:
"\<not>list_ex (is_LabelN l) A \<Longrightarrow> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \<subseteq> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l' A)"
by (rule trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(2)[of l A l']])
lemma setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label:
"\<not>list_ex (is_LabelN l) A \<Longrightarrow> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \<subseteq> setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l' A)"
by (rule setops\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(2)[of l A l']])
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_subset_if_no_label:
"\<not>list_ex (is_LabelN l) A \<Longrightarrow> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \<subseteq> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l' A)"
by (rule setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_subset_if_no_label(1)[of l A l']])
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,send\<langle>t\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,receive\<langle>t\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\<langle>ac: s \<doteq> t\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,insert\<langle>t,s\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {(l,t \<cdot> \<delta>,s \<cdot> \<delta>)}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,delete\<langle>t,s\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {(l,t \<cdot> \<delta>,s \<cdot> \<delta>)}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\<langle>ac: t \<in> s\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) = {(l,t \<cdot> \<delta>,s \<cdot> \<delta>)}"
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: F'\<rangle>) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>) =
((\<lambda>(t,s). (l,t \<cdot> rm_vars (set X) \<delta>,s \<cdot> rm_vars (set X) \<delta>)) ` set F')" (is "?A = ?B")
proof -
have "?A = (\<lambda>(t,s). (l,t,s)) ` set (F' \<cdot>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \<delta>)" by auto
thus "?A = ?B" unfolding subst_apply_pairs_def by auto
qed simp_all
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst:
assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \<inter> subst_domain \<theta> = {}"
shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>) = (\<lambda>p. (fst a,snd p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a"
proof -
obtain l a' where a: "a = (l,a')" by (metis surj_pair)
show ?thesis
proof (cases a')
case (NegChecks X F G)
hence *: "rm_vars (set X) \<theta> = \<theta>" using a assms rm_vars_apply'[of \<theta> "set X"] by auto
have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>) = (\<lambda>p. (fst a, p)) ` set (G \<cdot>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \<theta>)"
using * NegChecks a by auto
moreover have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\<lambda>p. (fst a, p)) ` set G" using NegChecks a by simp
hence "(\<lambda>p. (fst a,snd p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\<lambda>p. (fst a, p \<cdot>\<^sub>p \<theta>)) ` set G"
by (metis (mono_tags, lifting) image_cong image_image snd_conv)
hence "(\<lambda>p. (fst a,snd p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = (\<lambda>p. (fst a, p)) ` (set G \<cdot>\<^sub>p\<^sub>s\<^sub>e\<^sub>t \<theta>)"
unfolding case_prod_unfold by auto
ultimately show ?thesis by (simp add: subst_apply_pairs_def)
qed (use a in simp_all)
qed
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst':
assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \<inter> subst_domain \<theta> = {}"
shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>) = (\<lambda>(i,p). (i,p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a"
using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF assms] setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>pD(1) unfolding case_prod_unfold
by (metis (mono_tags, lifting) image_cong)
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst:
assumes "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<inter> subst_domain \<theta> = {}"
shows "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = (\<lambda>p. (fst p,snd p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S"
using assms
proof (induction S)
case (Cons a S)
have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \<inter> subst_domain \<theta> = {}" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \<inter> subst_domain \<theta> = {}"
using Cons.prems by auto
hence IH: "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = (\<lambda>p. (fst p,snd p \<cdot>\<^sub>p \<theta>)) ` setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S"
using Cons.IH by auto
show ?case
using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'[OF *] IH
unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold subst_lsst_cons
by auto
qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def)
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst:
assumes p: "p \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)"
shows "\<exists>q \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. fst p = fst q \<and> snd p = snd q \<cdot>\<^sub>p rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a))) \<delta>"
(is "\<exists>q \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. ?P q")
proof -
obtain l b where a: "a = (l,b)" by (metis surj_pair)
show ?thesis
proof (cases b)
case (NegChecks X F F')
hence "p \<in> (\<lambda>(t, s). (l, t \<cdot> rm_vars (set X) \<delta>, s \<cdot> rm_vars (set X) \<delta>)) ` set F'"
using p a setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of l X F F' \<delta>] by blast
then obtain s t where st:
"(t,s) \<in> set F'" "p = (l, t \<cdot> rm_vars (set X) \<delta>, s \<cdot> rm_vars (set X) \<delta>)"
by auto
hence "(l,t,s) \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst p = fst (l,t,s)"
"snd p = snd (l,t,s) \<cdot>\<^sub>p rm_vars (set X) \<delta>"
using a NegChecks by fastforce+
moreover have "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a) = X" using NegChecks a by auto
ultimately show ?thesis by blast
qed (use p a in auto)
qed
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst:
assumes "p \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)"
shows "\<exists>q \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. fst p = fst q \<and> (\<exists>X \<subseteq> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. snd p = snd q \<cdot>\<^sub>p rm_vars X \<delta>)"
(is "\<exists>q \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. ?P A q")
using assms
proof (induction A)
case (Cons a A)
note 0 = unlabel_Cons(2)[of a A] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of "snd a" "unlabel A"]
show ?case
proof (cases "p \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>)")
case False
hence "p \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>)"
using Cons.prems setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons[of "a \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<delta>" "A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>"] subst_lsst_cons[of a A \<delta>] by auto
moreover have "(set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a))) \<subseteq> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" using 0 by simp
ultimately have "\<exists>q \<in> setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. ?P (a#A) q" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst[of p a \<delta>] by blast
thus ?thesis by auto
qed (use Cons.IH 0 in auto)
qed simp
lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq:
"setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"
proof (induction A)
case (Cons a A)
obtain l b where "a = (l,b)" by (metis surj_pair)
thus ?case using Cons unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases b) auto
qed simp
end

View File

@ -0,0 +1,372 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2018-2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2018-2020
(C) Copyright Achim D. Brucker, University of Sheffield, 2018-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Labeled_Strands.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, The University of Sheffield
*)
section \<open>Labeled Strands\<close>
theory Labeled_Strands
imports Strands_and_Constraints
begin
subsection \<open>Definitions: Labeled Strands and Constraints\<close>
datatype 'l strand_label =
LabelN (the_LabelN: "'l") ("ln _")
| LabelS ("\<star>")
text \<open>Labeled strands are strands whose steps are equipped with labels\<close>
type_synonym ('a,'b,'c) labeled_strand_step = "'c strand_label \<times> ('a,'b) strand_step"
type_synonym ('a,'b,'c) labeled_strand = "('a,'b,'c) labeled_strand_step list"
abbreviation is_LabelN where "is_LabelN n x \<equiv> fst x = ln n"
abbreviation is_LabelS where "is_LabelS x \<equiv> fst x = \<star>"
definition unlabel where "unlabel S \<equiv> map snd S"
definition proj where "proj n S \<equiv> filter (\<lambda>s. is_LabelN n s \<or> is_LabelS s) S"
abbreviation proj_unl where "proj_unl n S \<equiv> unlabel (proj n S)"
abbreviation wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t where "wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t S \<equiv> wfrestrictedvars\<^sub>s\<^sub>t (unlabel S)"
abbreviation subst_apply_labeled_strand_step (infix "\<cdot>\<^sub>l\<^sub>s\<^sub>t\<^sub>p" 51) where
"x \<cdot>\<^sub>l\<^sub>s\<^sub>t\<^sub>p \<theta> \<equiv> (case x of (l, s) \<Rightarrow> (l, s \<cdot>\<^sub>s\<^sub>t\<^sub>p \<theta>))"
abbreviation subst_apply_labeled_strand (infix "\<cdot>\<^sub>l\<^sub>s\<^sub>t" 51) where
"S \<cdot>\<^sub>l\<^sub>s\<^sub>t \<theta> \<equiv> map (\<lambda>x. x \<cdot>\<^sub>l\<^sub>s\<^sub>t\<^sub>p \<theta>) S"
abbreviation trms\<^sub>l\<^sub>s\<^sub>t where "trms\<^sub>l\<^sub>s\<^sub>t S \<equiv> trms\<^sub>s\<^sub>t (unlabel S)"
abbreviation trms_proj\<^sub>l\<^sub>s\<^sub>t where "trms_proj\<^sub>l\<^sub>s\<^sub>t n S \<equiv> trms\<^sub>s\<^sub>t (proj_unl n S)"
abbreviation vars\<^sub>l\<^sub>s\<^sub>t where "vars\<^sub>l\<^sub>s\<^sub>t S \<equiv> vars\<^sub>s\<^sub>t (unlabel S)"
abbreviation vars_proj\<^sub>l\<^sub>s\<^sub>t where "vars_proj\<^sub>l\<^sub>s\<^sub>t n S \<equiv> vars\<^sub>s\<^sub>t (proj_unl n S)"
abbreviation bvars\<^sub>l\<^sub>s\<^sub>t where "bvars\<^sub>l\<^sub>s\<^sub>t S \<equiv> bvars\<^sub>s\<^sub>t (unlabel S)"
abbreviation fv\<^sub>l\<^sub>s\<^sub>t where "fv\<^sub>l\<^sub>s\<^sub>t S \<equiv> fv\<^sub>s\<^sub>t (unlabel S)"
abbreviation wf\<^sub>l\<^sub>s\<^sub>t where "wf\<^sub>l\<^sub>s\<^sub>t V S \<equiv> wf\<^sub>s\<^sub>t V (unlabel S)"
subsection \<open>Lemmata: Projections\<close>
lemma is_LabelS_proj_iff_not_is_LabelN:
"list_all is_LabelS (proj l A) \<longleftrightarrow> \<not>list_ex (is_LabelN l) A"
by (induct A) (auto simp add: proj_def)
lemma proj_subset_if_no_label:
assumes "\<not>list_ex (is_LabelN l) A"
shows "set (proj l A) \<subseteq> set (proj l' A)"
and "set (proj_unl l A) \<subseteq> set (proj_unl l' A)"
using assms by (induct A) (auto simp add: unlabel_def proj_def)
lemma proj_in_setD:
assumes a: "a \<in> set (proj l A)"
obtains k b where "a = (k, b)" "k = (ln l) \<or> k = \<star>"
using that a unfolding proj_def by (cases a) auto
lemma proj_set_mono:
assumes "set A \<subseteq> set B"
shows "set (proj n A) \<subseteq> set (proj n B)"
and "set (proj_unl n A) \<subseteq> set (proj_unl n B)"
using assms unfolding proj_def unlabel_def by auto
lemma unlabel_nil[simp]: "unlabel [] = []"
by (simp add: unlabel_def)
lemma unlabel_mono: "set A \<subseteq> set B \<Longrightarrow> set (unlabel A) \<subseteq> set (unlabel B)"
by (auto simp add: unlabel_def)
lemma unlabel_in: "(l,x) \<in> set A \<Longrightarrow> x \<in> set (unlabel A)"
unfolding unlabel_def by force
lemma unlabel_mem_has_label: "x \<in> set (unlabel A) \<Longrightarrow> \<exists>l. (l,x) \<in> set A"
unfolding unlabel_def by auto
lemma proj_nil[simp]: "proj n [] = []" "proj_unl n [] = []"
unfolding unlabel_def proj_def by auto
lemma singleton_lst_proj[simp]:
"proj_unl l [(ln l, a)] = [a]"
"l \<noteq> l' \<Longrightarrow> proj_unl l' [(ln l, a)] = []"
"proj_unl l [(\<star>, a)] = [a]"
"unlabel [(l'', a)] = [a]"
unfolding proj_def unlabel_def by simp_all
lemma unlabel_nil_only_if_nil[simp]: "unlabel A = [] \<Longrightarrow> A = []"
unfolding unlabel_def by auto
lemma unlabel_Cons[simp]:
"unlabel ((l,a)#A) = a#unlabel A"
"unlabel (b#A) = snd b#unlabel A"
unfolding unlabel_def by simp_all
lemma unlabel_append[simp]: "unlabel (A@B) = unlabel A@unlabel B"
unfolding unlabel_def by auto
lemma proj_Cons[simp]:
"proj n ((ln n,a)#A) = (ln n,a)#proj n A"
"proj n ((\<star>,a)#A) = (\<star>,a)#proj n A"
"m \<noteq> n \<Longrightarrow> proj n ((ln m,a)#A) = proj n A"
"l = (ln n) \<Longrightarrow> proj n ((l,a)#A) = (l,a)#proj n A"
"l = \<star> \<Longrightarrow> proj n ((l,a)#A) = (l,a)#proj n A"
"fst b \<noteq> \<star> \<Longrightarrow> fst b \<noteq> (ln n) \<Longrightarrow> proj n (b#A) = proj n A"
unfolding proj_def by auto
lemma proj_append[simp]:
"proj l (A'@B') = proj l A'@proj l B'"
"proj_unl l (A@B) = proj_unl l A@proj_unl l B"
unfolding proj_def unlabel_def by auto
lemma proj_unl_cons[simp]:
"proj_unl l ((ln l, a)#A) = a#proj_unl l A"
"l \<noteq> l' \<Longrightarrow> proj_unl l' ((ln l, a)#A) = proj_unl l' A"
"proj_unl l ((\<star>, a)#A) = a#proj_unl l A"
unfolding proj_def unlabel_def by simp_all
lemma trms_unlabel_proj[simp]:
"trms\<^sub>s\<^sub>t\<^sub>p (snd (ln l, x)) \<subseteq> trms_proj\<^sub>l\<^sub>s\<^sub>t l [(ln l, x)]"
by auto
lemma trms_unlabel_star[simp]:
"trms\<^sub>s\<^sub>t\<^sub>p (snd (\<star>, x)) \<subseteq> trms_proj\<^sub>l\<^sub>s\<^sub>t l [(\<star>, x)]"
by auto
lemma trms\<^sub>l\<^sub>s\<^sub>t_union[simp]: "trms\<^sub>l\<^sub>s\<^sub>t A = (\<Union>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A)"
proof (induction A)
case (Cons a A)
obtain l s where ls: "a = (l,s)" by moura
have "trms\<^sub>l\<^sub>s\<^sub>t [a] = (\<Union>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a])"
proof -
have *: "trms\<^sub>l\<^sub>s\<^sub>t [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by simp
show ?thesis
proof (cases l)
case (LabelN n)
hence "trms_proj\<^sub>l\<^sub>s\<^sub>t n [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by simp
moreover have "\<forall>m. n \<noteq> m \<longrightarrow> trms_proj\<^sub>l\<^sub>s\<^sub>t m [a] = {}" using ls LabelN by auto
ultimately show ?thesis using * ls by fastforce
next
case LabelS
hence "\<forall>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a] = trms\<^sub>s\<^sub>t\<^sub>p s" using ls by auto
thus ?thesis using * ls by fastforce
qed
qed
moreover have "\<forall>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l (a#A) = trms_proj\<^sub>l\<^sub>s\<^sub>t l [a] \<union> trms_proj\<^sub>l\<^sub>s\<^sub>t l A"
unfolding unlabel_def proj_def by auto
hence "(\<Union>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l (a#A)) = (\<Union>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l [a]) \<union> (\<Union>l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A)" by auto
ultimately show ?case using Cons.IH ls by auto
qed simp
lemma trms\<^sub>l\<^sub>s\<^sub>t_append[simp]: "trms\<^sub>l\<^sub>s\<^sub>t (A@B) = trms\<^sub>l\<^sub>s\<^sub>t A \<union> trms\<^sub>l\<^sub>s\<^sub>t B"
by (metis trms\<^sub>s\<^sub>t_append unlabel_append)
lemma trms_proj\<^sub>l\<^sub>s\<^sub>t_append[simp]: "trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B) = trms_proj\<^sub>l\<^sub>s\<^sub>t l A \<union> trms_proj\<^sub>l\<^sub>s\<^sub>t l B"
by (metis (no_types, lifting) filter_append proj_def trms\<^sub>l\<^sub>s\<^sub>t_append)
lemma trms_proj\<^sub>l\<^sub>s\<^sub>t_subset[simp]:
"trms_proj\<^sub>l\<^sub>s\<^sub>t l A \<subseteq> trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B)"
"trms_proj\<^sub>l\<^sub>s\<^sub>t l B \<subseteq> trms_proj\<^sub>l\<^sub>s\<^sub>t l (A@B)"
using trms_proj\<^sub>l\<^sub>s\<^sub>t_append[of l] by blast+
lemma trms\<^sub>l\<^sub>s\<^sub>t_subset[simp]:
"trms\<^sub>l\<^sub>s\<^sub>t A \<subseteq> trms\<^sub>l\<^sub>s\<^sub>t (A@B)"
"trms\<^sub>l\<^sub>s\<^sub>t B \<subseteq> trms\<^sub>l\<^sub>s\<^sub>t (A@B)"
proof (induction A)
case (Cons a A)
obtain l s where *: "a = (l,s)" by moura
{ case 1 thus ?case using Cons * by auto }
{ case 2 thus ?case using Cons * by auto }
qed simp_all
lemma vars\<^sub>l\<^sub>s\<^sub>t_union: "vars\<^sub>l\<^sub>s\<^sub>t A = (\<Union>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l A)"
proof (induction A)
case (Cons a A)
obtain l s where ls: "a = (l,s)" by moura
have "vars\<^sub>l\<^sub>s\<^sub>t [a] = (\<Union>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a])"
proof -
have *: "vars\<^sub>l\<^sub>s\<^sub>t [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by auto
show ?thesis
proof (cases l)
case (LabelN n)
hence "vars_proj\<^sub>l\<^sub>s\<^sub>t n [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by simp
moreover have "\<forall>m. n \<noteq> m \<longrightarrow> vars_proj\<^sub>l\<^sub>s\<^sub>t m [a] = {}" using ls LabelN by auto
ultimately show ?thesis using * ls by fast
next
case LabelS
hence "\<forall>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a] = vars\<^sub>s\<^sub>t\<^sub>p s" using ls by auto
thus ?thesis using * ls by fast
qed
qed
moreover have "\<forall>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l (a#A) = vars_proj\<^sub>l\<^sub>s\<^sub>t l [a] \<union> vars_proj\<^sub>l\<^sub>s\<^sub>t l A"
unfolding unlabel_def proj_def by auto
hence "(\<Union>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l (a#A)) = (\<Union>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l [a]) \<union> (\<Union>l. vars_proj\<^sub>l\<^sub>s\<^sub>t l A)"
using strand_vars_split(1) by auto
ultimately show ?case using Cons.IH ls strand_vars_split(1) by auto
qed simp
lemma unlabel_Cons_inv:
"unlabel A = b#B \<Longrightarrow> \<exists>A'. (\<exists>n. A = (ln n, b)#A') \<or> A = (\<star>, b)#A'"
proof -
assume *: "unlabel A = b#B"
then obtain l A' where "A = (l,b)#A'" unfolding unlabel_def by moura
thus "\<exists>A'. (\<exists>l. A = (ln l, b)#A') \<or> A = (\<star>, b)#A'" by (metis strand_label.exhaust)
qed
lemma unlabel_snoc_inv:
"unlabel A = B@[b] \<Longrightarrow> \<exists>A'. (\<exists>n. A = A'@[(ln n, b)]) \<or> A = A'@[(\<star>, b)]"
proof -
assume *: "unlabel A = B@[b]"
then obtain A' l where "A = A'@[(l,b)]"
unfolding unlabel_def by (induct A rule: List.rev_induct) auto
thus "\<exists>A'. (\<exists>n. A = A'@[(ln n, b)]) \<or> A = A'@[(\<star>, b)]" by (cases l) auto
qed
lemma proj_idem[simp]: "proj l (proj l A) = proj l A"
unfolding proj_def by auto
lemma proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set:
"ik\<^sub>s\<^sub>t (proj_unl n A) = {t. (ln n, Receive t) \<in> set A \<or> (\<star>, Receive t) \<in> set A} "
using ik\<^sub>s\<^sub>t_is_rcv_set unfolding unlabel_def proj_def by force
lemma unlabel_ik\<^sub>s\<^sub>t_is_rcv_set:
"ik\<^sub>s\<^sub>t (unlabel A) = {t | l t. (l, Receive t) \<in> set A}"
using ik\<^sub>s\<^sub>t_is_rcv_set unfolding unlabel_def by force
lemma proj_ik_union_is_unlabel_ik:
"ik\<^sub>s\<^sub>t (unlabel A) = (\<Union>l. ik\<^sub>s\<^sub>t (proj_unl l A))"
proof
show "(\<Union>l. ik\<^sub>s\<^sub>t (proj_unl l A)) \<subseteq> ik\<^sub>s\<^sub>t (unlabel A)"
using unlabel_ik\<^sub>s\<^sub>t_is_rcv_set[of A] proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set[of _ A] by auto
show "ik\<^sub>s\<^sub>t (unlabel A) \<subseteq> (\<Union>l. ik\<^sub>s\<^sub>t (proj_unl l A))"
proof
fix t assume "t \<in> ik\<^sub>s\<^sub>t (unlabel A)"
then obtain l where "(l, Receive t) \<in> set A"
using ik\<^sub>s\<^sub>t_is_rcv_set unlabel_mem_has_label[of _ A]
by moura
thus "t \<in> (\<Union>l. ik\<^sub>s\<^sub>t (proj_unl l A))" using proj_ik\<^sub>s\<^sub>t_is_proj_rcv_set[of _ A] by (cases l) auto
qed
qed
lemma proj_ik_append[simp]:
"ik\<^sub>s\<^sub>t (proj_unl l (A@B)) = ik\<^sub>s\<^sub>t (proj_unl l A) \<union> ik\<^sub>s\<^sub>t (proj_unl l B)"
using proj_append(2)[of l A B] ik_append by auto
lemma proj_ik_append_subst_all:
"ik\<^sub>s\<^sub>t (proj_unl l (A@B)) \<cdot>\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>s\<^sub>t (proj_unl l A) \<cdot>\<^sub>s\<^sub>e\<^sub>t I) \<union> (ik\<^sub>s\<^sub>t (proj_unl l B) \<cdot>\<^sub>s\<^sub>e\<^sub>t I)"
using proj_ik_append[of l] by auto
lemma ik_proj_subset[simp]: "ik\<^sub>s\<^sub>t (proj_unl n A) \<subseteq> trms_proj\<^sub>l\<^sub>s\<^sub>t n A"
by auto
lemma prefix_proj:
"prefix A B \<Longrightarrow> prefix (unlabel A) (unlabel B)"
"prefix A B \<Longrightarrow> prefix (proj n A) (proj n B)"
"prefix A B \<Longrightarrow> prefix (proj_unl n A) (proj_unl n B)"
unfolding prefix_def unlabel_def proj_def by auto
subsection \<open>Lemmata: Well-formedness\<close>
lemma wfvarsoccs\<^sub>s\<^sub>t_proj_union:
"wfvarsoccs\<^sub>s\<^sub>t (unlabel A) = (\<Union>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A))"
proof (induction A)
case (Cons a A)
obtain l s where ls: "a = (l,s)" by moura
have "wfvarsoccs\<^sub>s\<^sub>t (unlabel [a]) = (\<Union>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]))"
proof -
have *: "wfvarsoccs\<^sub>s\<^sub>t (unlabel [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by auto
show ?thesis
proof (cases l)
case (LabelN n)
hence "wfvarsoccs\<^sub>s\<^sub>t (proj_unl n [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by simp
moreover have "\<forall>m. n \<noteq> m \<longrightarrow> wfvarsoccs\<^sub>s\<^sub>t (proj_unl m [a]) = {}" using ls LabelN by auto
ultimately show ?thesis using * ls by fast
next
case LabelS
hence "\<forall>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]) = wfvarsoccs\<^sub>s\<^sub>t\<^sub>p s" using ls by auto
thus ?thesis using * ls by fast
qed
qed
moreover have
"wfvarsoccs\<^sub>s\<^sub>t (proj_unl l (a#A)) =
wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a]) \<union> wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)"
for l
unfolding unlabel_def proj_def by auto
hence "(\<Union>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l (a#A))) =
(\<Union>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l [a])) \<union> (\<Union>l. wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A))"
using strand_vars_split(1) by auto
ultimately show ?case using Cons.IH ls strand_vars_split(1) by auto
qed simp
lemma wf_if_wf_proj:
assumes "\<forall>l. wf\<^sub>s\<^sub>t V (proj_unl l A)"
shows "wf\<^sub>s\<^sub>t V (unlabel A)"
using assms
proof (induction A arbitrary: V rule: List.rev_induct)
case (snoc a A)
hence IH: "wf\<^sub>s\<^sub>t V (unlabel A)" using proj_append(2)[of _ A] by auto
obtain b l where b: "a = (ln l, b) \<or> a = (\<star>, b)" by (cases a, metis strand_label.exhaust)
hence *: "wf\<^sub>s\<^sub>t V (proj_unl l A@[b])"
by (metis snoc.prems proj_append(2) singleton_lst_proj(1) proj_unl_cons(1,3))
thus ?case using IH b snoc.prems proj_append(2)[of l A "[a]"] unlabel_append[of A "[a]"]
proof (cases b)
case (Receive t)
have "fv t \<subseteq> wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \<union> V"
proof
fix x assume "x \<in> fv t"
hence "x \<in> V \<union> wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)" using wf_append_exec[OF *] b Receive by auto
thus "x \<in> wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \<union> V" using wfvarsoccs\<^sub>s\<^sub>t_proj_union[of A] by auto
qed
hence "fv t \<subseteq> wfrestrictedvars\<^sub>s\<^sub>t (unlabel A) \<union> V"
using vars_snd_rcv_strand_subset2(4)[of "unlabel A"] by blast
hence "wf\<^sub>s\<^sub>t V (unlabel A@[Receive t])" by (rule wf_rcv_append'''[OF IH])
thus ?thesis using b Receive unlabel_append[of A "[a]"] by auto
next
case (Equality ac s t)
have "fv t \<subseteq> wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \<union> V" when "ac = Assign"
proof
fix x assume "x \<in> fv t"
hence "x \<in> V \<union> wfvarsoccs\<^sub>s\<^sub>t (proj_unl l A)" using wf_append_exec[OF *] b Equality that by auto
thus "x \<in> wfvarsoccs\<^sub>s\<^sub>t (unlabel A) \<union> V" using wfvarsoccs\<^sub>s\<^sub>t_proj_union[of A] by auto
qed
hence "fv t \<subseteq> wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t A \<union> V" when "ac = Assign"
using vars_snd_rcv_strand_subset2(4)[of "unlabel A"] that by blast
hence "wf\<^sub>s\<^sub>t V (unlabel A@[Equality ac s t])"
by (cases ac) (metis wf_eq_append'''[OF IH], metis wf_eq_check_append''[OF IH])
thus ?thesis using b Equality unlabel_append[of A "[a]"] by auto
qed auto
qed simp
end

View File

@ -0,0 +1,884 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Lazy_Intruder.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>The Lazy Intruder\<close>
theory Lazy_Intruder
imports Strands_and_Constraints Intruder_Deduction
begin
context intruder_model
begin
subsection \<open>Definition of the Lazy Intruder\<close>
text \<open>The lazy intruder constraint reduction system, defined as a relation on constraint states\<close>
inductive_set LI_rel::
"((('fun,'var) strand \<times> (('fun,'var) subst)) \<times>
('fun,'var) strand \<times> (('fun,'var) subst)) set"
and LI_rel' (infix "\<leadsto>" 50)
and LI_rel_trancl (infix "\<leadsto>\<^sup>+" 50)
and LI_rel_rtrancl (infix "\<leadsto>\<^sup>*" 50)
where
"A \<leadsto> B \<equiv> (A,B) \<in> LI_rel"
| "A \<leadsto>\<^sup>+ B \<equiv> (A,B) \<in> LI_rel\<^sup>+"
| "A \<leadsto>\<^sup>* B \<equiv> (A,B) \<in> LI_rel\<^sup>*"
| Compose: "\<lbrakk>simple S; length T = arity f; public f\<rbrakk>
\<Longrightarrow> (S@Send (Fun f T)#S',\<theta>) \<leadsto> (S@(map Send T)@S',\<theta>)"
| Unify: "\<lbrakk>simple S; Fun f T' \<in> ik\<^sub>s\<^sub>t S; Some \<delta> = mgu (Fun f T) (Fun f T')\<rbrakk>
\<Longrightarrow> (S@Send (Fun f T)#S',\<theta>) \<leadsto> ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta> \<circ>\<^sub>s \<delta>)"
| Equality: "\<lbrakk>simple S; Some \<delta> = mgu t t'\<rbrakk>
\<Longrightarrow> (S@Equality _ t t'#S',\<theta>) \<leadsto> ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta> \<circ>\<^sub>s \<delta>)"
subsection \<open>Lemma: The Lazy Intruder is Well-founded\<close>
context
begin
private lemma LI_compose_measure_lt: "((S@(map Send T)@S',\<theta>\<^sub>1), (S@Send (Fun f T)#S',\<theta>\<^sub>2)) \<in> measure\<^sub>s\<^sub>t"
using strand_fv_card_map_fun_eq[of S f T S'] strand_size_map_fun_lt(2)[of T f]
by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
private lemma LI_unify_measure_lt:
assumes "Some \<delta> = mgu (Fun f T) t" "fv t \<subseteq> fv\<^sub>s\<^sub>t S"
shows "(((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta>\<^sub>1), (S@Send (Fun f T)#S',\<theta>\<^sub>2)) \<in> measure\<^sub>s\<^sub>t"
proof (cases "\<delta> = Var")
assume "\<delta> = Var"
hence "(S@S') \<cdot>\<^sub>s\<^sub>t \<delta> = S@S'" by blast
thus ?thesis
using strand_fv_card_rm_fun_le[of S S' f T]
by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
next
assume "\<delta> \<noteq> Var"
then obtain v where "v \<in> fv (Fun f T) \<union> fv t" "subst_elim \<delta> v"
using mgu_eliminates[OF assms(1)[symmetric]] by metis
hence v_in: "v \<in> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using assms(2) by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
have "range_vars \<delta> \<subseteq> fv (Fun f T) \<union> fv\<^sub>s\<^sub>t S"
using assms(2) mgu_vars_bounded[OF assms(1)[symmetric]] by auto
hence img_bound: "range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" by auto
have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S'))" by auto
have "v \<notin> fv\<^sub>s\<^sub>t ((S@Send (Fun f T)#S') \<cdot>\<^sub>s\<^sub>t \<delta>)"
using strand_fv_subst_subset_if_subst_elim[OF \<open>subst_elim \<delta> v\<close>] v_in by metis
hence v_not_in: "v \<notin> fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>)" by auto
have "fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp
hence "fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>) \<subset> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using v_in v_not_in by blast
hence "card (fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>)) < card (fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S'))"
using psubset_card_mono[OF finite_fv] by simp
thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
qed
private lemma LI_equality_measure_lt:
assumes "Some \<delta> = mgu t t'"
shows "(((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta>\<^sub>1), (S@Equality a t t'#S',\<theta>\<^sub>2)) \<in> measure\<^sub>s\<^sub>t"
proof (cases "\<delta> = Var")
assume "\<delta> = Var"
hence "(S@S') \<cdot>\<^sub>s\<^sub>t \<delta> = S@S'" by blast
thus ?thesis
using strand_fv_card_rm_eq_le[of S S' a t t']
by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
next
assume "\<delta> \<noteq> Var"
then obtain v where "v \<in> fv t \<union> fv t'" "subst_elim \<delta> v"
using mgu_eliminates[OF assms(1)[symmetric]] by metis
hence v_in: "v \<in> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using assms by auto
have "range_vars \<delta> \<subseteq> fv t \<union> fv t' \<union> fv\<^sub>s\<^sub>t S"
using assms mgu_vars_bounded[OF assms(1)[symmetric]] by auto
hence img_bound: "range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto
have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by auto
have "v \<notin> fv\<^sub>s\<^sub>t ((S@Equality a t t'#S') \<cdot>\<^sub>s\<^sub>t \<delta>)"
using strand_fv_subst_subset_if_subst_elim[OF \<open>subst_elim \<delta> v\<close>] v_in by metis
hence v_not_in: "v \<notin> fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>)" by auto
have "fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp
hence "fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>) \<subset> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using v_in v_not_in by blast
hence "card (fv\<^sub>s\<^sub>t ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>)) < card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))"
using psubset_card_mono[OF finite_fv] by simp
thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def)
qed
private lemma LI_in_measure: "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto> (S\<^sub>2,\<theta>\<^sub>2) \<Longrightarrow> ((S\<^sub>2,\<theta>\<^sub>2),(S\<^sub>1,\<theta>\<^sub>1)) \<in> measure\<^sub>s\<^sub>t"
proof (induction rule: LI_rel.induct)
case (Compose S T f S' \<theta>) thus ?case using LI_compose_measure_lt[of S T S'] by metis
next
case (Unify S f U \<delta> T S' \<theta>)
hence "fv (Fun f U) \<subseteq> fv\<^sub>s\<^sub>t S"
using fv_snd_rcv_strand_subset(2)[of S] by force
thus ?case using LI_unify_measure_lt[OF Unify.hyps(3), of S S'] by metis
qed (metis LI_equality_measure_lt)
private lemma LI_in_measure_trans: "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>+ (S\<^sub>2,\<theta>\<^sub>2) \<Longrightarrow> ((S\<^sub>2,\<theta>\<^sub>2),(S\<^sub>1,\<theta>\<^sub>1)) \<in> measure\<^sub>s\<^sub>t"
by (induction rule: trancl.induct, metis surjective_pairing LI_in_measure)
(metis (no_types, lifting) surjective_pairing LI_in_measure measure\<^sub>s\<^sub>t_trans trans_def)
private lemma LI_converse_wellfounded_trans: "wf ((LI_rel\<^sup>+)\<inverse>)"
proof -
have "(LI_rel\<^sup>+)\<inverse> \<subseteq> measure\<^sub>s\<^sub>t" using LI_in_measure_trans by auto
thus ?thesis using measure\<^sub>s\<^sub>t_wellfounded wf_subset by metis
qed
private lemma LI_acyclic_trans: "acyclic (LI_rel\<^sup>+)"
using wf_acyclic[OF LI_converse_wellfounded_trans] acyclic_converse by metis
private lemma LI_acyclic: "acyclic LI_rel"
using LI_acyclic_trans acyclic_subset by (simp add: acyclic_def)
lemma LI_no_infinite_chain: "\<not>(\<exists>f. \<forall>i. f i \<leadsto>\<^sup>+ f (Suc i))"
proof -
have "\<not>(\<exists>f. \<forall>i. (f (Suc i), f i) \<in> (LI_rel\<^sup>+)\<inverse>)"
using wf_iff_no_infinite_down_chain LI_converse_wellfounded_trans by metis
thus ?thesis by simp
qed
private lemma LI_unify_finite:
assumes "finite M"
shows "finite {((S@Send (Fun f T)#S',\<theta>), ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta> \<circ>\<^sub>s \<delta>)) | \<delta> T'.
simple S \<and> Fun f T' \<in> M \<and> Some \<delta> = mgu (Fun f T) (Fun f T')}"
using assms
proof (induction M rule: finite_induct)
case (insert m M) thus ?case
proof (cases m)
case (Fun g U)
let ?a = "\<lambda>\<delta>. ((S@Send (Fun f T)#S',\<theta>), ((S@S') \<cdot>\<^sub>s\<^sub>t \<delta>,\<theta> \<circ>\<^sub>s \<delta>))"
let ?A = "\<lambda>B. {?a \<delta> | \<delta> T'. simple S \<and> Fun f T' \<in> B \<and> Some \<delta> = mgu (Fun f T) (Fun f T')}"
have "?A (insert m M) = (?A M) \<union> (?A {m})" by auto
moreover have "finite (?A {m})"
proof (cases "\<exists>\<delta>. Some \<delta> = mgu (Fun f T) (Fun g U)")
case True
then obtain \<delta> where \<delta>: "Some \<delta> = mgu (Fun f T) (Fun g U)" by blast
have A_m_eq: "\<And>\<delta>'. ?a \<delta>' \<in> ?A {m} \<Longrightarrow> ?a \<delta> = ?a \<delta>'"
proof -
fix \<delta>' assume "?a \<delta>' \<in> ?A {m}"
hence "\<exists>\<sigma>. Some \<sigma> = mgu (Fun f T) (Fun g U) \<and> ?a \<sigma> = ?a \<delta>'"
using \<open>m = Fun g U\<close> by auto
thus "?a \<delta> = ?a \<delta>'" by (metis \<delta> option.inject)
qed
have "?A {m} = {} \<or> ?A {m} = {?a \<delta>}"
proof (cases "simple S \<and> ?A {m} \<noteq> {}")
case True
hence "simple S" "?A {m} \<noteq> {}" by meson+
hence "?A {m} = {?a \<delta> | \<delta>. Some \<delta> = mgu (Fun f T) (Fun g U)}" using \<open>m = Fun g U\<close> by auto
hence "?a \<delta> \<in> ?A {m}" using \<delta> by auto
show ?thesis
proof (rule ccontr)
assume "\<not>(?A {m} = {} \<or> ?A {m} = {?a \<delta>})"
then obtain B where B: "?A {m} = insert (?a \<delta>) B" "?a \<delta> \<notin> B" "B \<noteq> {}"
using \<open>?A {m} \<noteq> {}\<close> \<open>?a \<delta> \<in> ?A {m}\<close> by (metis (no_types, lifting) Set.set_insert)
then obtain b where b: "?a \<delta> \<noteq> b" "b \<in> B" by (metis (no_types, lifting) ex_in_conv)
then obtain \<delta>' where \<delta>': "b = ?a \<delta>'" using B(1) by blast
moreover have "?a \<delta>' \<in> ?A {m}" using B(1) b(2) \<delta>' by auto
hence "?a \<delta> = ?a \<delta>'" by (blast dest!: A_m_eq)
ultimately show False using b(1) by simp
qed
qed auto
thus ?thesis by (metis (no_types, lifting) finite.emptyI finite_insert)
next
case False
hence "?A {m} = {}" using \<open>m = Fun g U\<close> by blast
thus ?thesis by (metis finite.emptyI)
qed
ultimately show ?thesis using insert.IH by auto
qed simp
qed fastforce
end
subsection \<open>Lemma: The Lazy Intruder Preserves Well-formedness\<close>
context
begin
private lemma LI_preserves_subst_wf_single:
assumes "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto> (S\<^sub>2,\<theta>\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<^sub>1"
and "subst_domain \<theta>\<^sub>1 \<inter> vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \<theta>\<^sub>1 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>1 = {}"
shows "fv\<^sub>s\<^sub>t S\<^sub>2 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<^sub>2"
and "subst_domain \<theta>\<^sub>2 \<inter> vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \<theta>\<^sub>2 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>2 = {}"
using assms
proof (induction rule: LI_rel.induct)
case (Compose S X f S' \<theta>)
{ case 1 thus ?case using vars_st_snd_map by auto }
{ case 2 thus ?case using vars_st_snd_map by auto }
{ case 3 thus ?case using vars_st_snd_map by force }
{ case 4 thus ?case using vars_st_snd_map by auto }
next
case (Unify S f U \<delta> T S' \<theta>)
hence "fv (Fun f U) \<subseteq> fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik' by blast
hence *: "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]]
unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps)
have "fv\<^sub>s\<^sub>t (S@S') \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" "vars\<^sub>s\<^sub>t (S@S') \<subseteq> vars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
by auto
hence **: "fv\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
"vars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> vars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using subst_sends_strand_fv_to_img[of "S@S'" \<delta>]
strand_subst_vars_union_bound[of "S@S'" \<delta>] *
by blast+
have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<delta>" by (fact mgu_gives_wellformed_subst[OF Unify.hyps(3)[symmetric]])
{ case 1
have "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
thus ?case using 1 ** by blast
}
{ case 2
hence "subst_domain \<theta> \<inter> subst_domain \<delta> = {}" "subst_domain \<theta> \<inter> range_vars \<delta> = {}"
using * by blast+
thus ?case by (metis wf_subst_compose[OF \<open>wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<close> \<open>wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<delta>\<close>])
}
{ case 3
hence "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using ** by blast
moreover have "v \<in> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" when "v \<in> subst_domain \<delta>" for v
using * that by blast
hence "subst_domain \<delta> \<inter> fv\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
using mgu_eliminates_dom[OF Unify.hyps(3)[symmetric],
THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Send (Fun f T)#S'"]
unfolding subst_elim_def by auto
moreover have "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
hence "subst_domain \<delta> \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using 3(1) * by blast
ultimately show ?case
using ** * subst_domain_compose[of \<theta> \<delta>] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \<cdot>\<^sub>s\<^sub>t \<delta>"]
by blast
}
{ case 4
have ***: "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
hence "range_vars \<delta> \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using 4(1) * by blast
thus ?case using subst_img_comp_subset[of \<theta> \<delta>] 4(4) *** by blast
}
next
case (Equality S \<delta> t t' a S' \<theta>)
hence *: "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]]
unfolding range_vars_alt_def by fastforce
have "fv\<^sub>s\<^sub>t (S@S') \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S') \<subseteq> vars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
by auto
hence **: "fv\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')"
"vars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<subseteq> vars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using subst_sends_strand_fv_to_img[of "S@S'" \<delta>]
strand_subst_vars_union_bound[of "S@S'" \<delta>] *
by blast+
have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<delta>" by (fact mgu_gives_wellformed_subst[OF Equality.hyps(2)[symmetric]])
{ case 1
have "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
thus ?case using 1 ** by blast
}
{ case 2
hence "subst_domain \<theta> \<inter> subst_domain \<delta> = {}" "subst_domain \<theta> \<inter> range_vars \<delta> = {}"
using * by blast+
thus ?case by (metis wf_subst_compose[OF \<open>wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<close> \<open>wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<delta>\<close>])
}
{ case 3
hence "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using ** by blast
moreover have "v \<in> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" when "v \<in> subst_domain \<delta>" for v
using * that by blast
hence "subst_domain \<delta> \<inter> fv\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
using mgu_eliminates_dom[OF Equality.hyps(2)[symmetric],
THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Equality a t t'#S'"]
unfolding subst_elim_def by auto
moreover have "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
hence "subst_domain \<delta> \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using 3(1) * by blast
ultimately show ?case
using ** * subst_domain_compose[of \<theta> \<delta>] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \<cdot>\<^sub>s\<^sub>t \<delta>"]
by blast
}
{ case 4
have ***: "bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using bvars_subst_ident[of "S@S'" \<delta>] by auto
hence "range_vars \<delta> \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}" using 4(1) * by blast
thus ?case using subst_img_comp_subset[of \<theta> \<delta>] 4(4) *** by blast
}
qed
private lemma LI_preserves_subst_wf:
assumes "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>2,\<theta>\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<^sub>1"
and "subst_domain \<theta>\<^sub>1 \<inter> vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \<theta>\<^sub>1 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>1 = {}"
shows "fv\<^sub>s\<^sub>t S\<^sub>2 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<theta>\<^sub>2"
and "subst_domain \<theta>\<^sub>2 \<inter> vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \<theta>\<^sub>2 \<inter> bvars\<^sub>s\<^sub>t S\<^sub>2 = {}"
using assms
proof (induction S\<^sub>2 \<theta>\<^sub>2 rule: rtrancl_induct2)
case (step S\<^sub>i \<theta>\<^sub>i S\<^sub>j \<theta>\<^sub>j)
{ case 1 thus ?case using LI_preserves_subst_wf_single[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>] step.IH by metis }
{ case 2 thus ?case using LI_preserves_subst_wf_single[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>] step.IH by metis }
{ case 3 thus ?case using LI_preserves_subst_wf_single[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>] step.IH by metis }
{ case 4 thus ?case using LI_preserves_subst_wf_single[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>] step.IH by metis }
qed metis
lemma LI_preserves_wellformedness:
assumes "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>2,\<theta>\<^sub>2)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1"
shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>2 \<theta>\<^sub>2"
proof -
have *: "wf\<^sub>s\<^sub>t {} S\<^sub>j"
when "(S\<^sub>i, \<theta>\<^sub>i) \<leadsto> (S\<^sub>j, \<theta>\<^sub>j)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>i \<theta>\<^sub>i" for S\<^sub>i \<theta>\<^sub>i S\<^sub>j \<theta>\<^sub>j
using that
proof (induction rule: LI_rel.induct)
case (Unify S f U \<delta> T S' \<theta>)
have "fv (Fun f T) \<union> fv (Fun f U) \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using Unify.hyps(2) by force
hence "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by (metis subset_trans)
hence "(subst_domain \<delta> \<union> range_vars \<delta>) \<inter> bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S') = {}"
using Unify.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast
thus ?case
using wf_unify[OF _ Unify.hyps(2) MGU_is_Unifier[OF mgu_gives_MGU], of "{}",
OF _ Unify.hyps(3)[symmetric], of S'] Unify.prems(1)
by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
next
case (Equality S \<delta> t t' a S' \<theta>)
have "fv t \<union> fv t' \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using Equality.hyps(2) by force
hence "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by (metis subset_trans)
hence "(subst_domain \<delta> \<union> range_vars \<delta>) \<inter> bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}"
using Equality.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast
thus ?case
using wf_equality[OF _ Equality.hyps(2)[symmetric], of "{}" S a S'] Equality.prems(1)
by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
qed (metis wf_send_compose wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
show ?thesis using assms
proof (induction rule: rtrancl_induct2)
case (step S\<^sub>i \<theta>\<^sub>i S\<^sub>j \<theta>\<^sub>j) thus ?case
using LI_preserves_subst_wf_single[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>] *[OF \<open>(S\<^sub>i,\<theta>\<^sub>i) \<leadsto> (S\<^sub>j,\<theta>\<^sub>j)\<close>]
by (metis wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
qed simp
qed
lemma LI_preserves_trm_wf:
assumes "(S,\<theta>) \<leadsto>\<^sup>* (S',\<theta>')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)"
shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')"
proof -
{ fix S \<theta> S' \<theta>'
assume "(S,\<theta>) \<leadsto> (S',\<theta>')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)"
hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')"
proof (induction rule: LI_rel.induct)
case (Compose S T f S' \<theta>)
hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)"
and *: "t \<in> set S \<Longrightarrow> wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" "t \<in> set S' \<Longrightarrow> wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" for t
by auto
hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \<in> set T" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto
hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" when "t \<in> set (map Send T)" for t
using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto
thus ?case using * by force
next
case (Unify S f U \<delta> T S' \<theta>)
have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f U)"
using Unify.prems(1) Unify.hyps(2) wf_trm_subterm[of _ "Fun f U"]
by (simp, force)
hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \<delta>)"
using mgu_wf_trm[OF Unify.hyps(3)[symmetric]] by simp
{ fix s assume "s \<in> set (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>)"
hence "\<exists>s' \<in> set (S@S'). s = s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta> \<and> wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')"
using Unify.prems(1) by (auto simp add: subst_apply_strand_def)
moreover {
fix s' assume s': "s = s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta>" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \<in> set (S@S')"
from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta>) = trms\<^sub>s\<^sub>t\<^sub>p s' \<cdot>\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \<delta>)"
proof (induction s')
case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def)
qed auto
hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)"
using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \<open>wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\<close> s'(1)
by simp
}
ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto
}
thus ?case by auto
next
case (Equality S \<delta> t t' a S' \<theta>)
hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" by simp_all
hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \<delta>)"
using mgu_wf_trm[OF Equality.hyps(2)[symmetric]] by simp
{ fix s assume "s \<in> set (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>)"
hence "\<exists>s' \<in> set (S@S'). s = s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta> \<and> wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')"
using Equality.prems(1) by (auto simp add: subst_apply_strand_def)
moreover {
fix s' assume s': "s = s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta>" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \<in> set (S@S')"
from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \<cdot>\<^sub>s\<^sub>t\<^sub>p \<delta>) = trms\<^sub>s\<^sub>t\<^sub>p s' \<cdot>\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \<delta>)"
proof (induction s')
case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def)
qed auto
hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)"
using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \<open>wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\<close> s'(1)
by simp
}
ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto
}
thus ?case by auto
qed
}
with assms show ?thesis by (induction rule: rtrancl_induct2) metis+
qed
end
subsection \<open>Theorem: Soundness of the Lazy Intruder\<close>
context
begin
private lemma LI_soundness_single:
assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1" "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto> (S\<^sub>2,\<theta>\<^sub>2)" "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2,\<theta>\<^sub>2\<rangle>"
shows "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1,\<theta>\<^sub>1\<rangle>"
using assms(2,1,3)
proof (induction rule: LI_rel.induct)
case (Compose S T f S' \<theta>)
hence *: "\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; map Send T\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>"
unfolding constr_sem_c_def by force+
have "ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c Fun f T \<cdot> \<I>"
using *(2) Compose.hyps(2) ComposeC[OF _ Compose.hyps(3), of "map (\<lambda>x. x \<cdot> \<I>) T"]
unfolding subst_compose_def by force
thus "\<I> \<Turnstile>\<^sub>c \<langle>S@Send (Fun f T)#S',\<theta>\<rangle>"
using *(1,3) \<open>\<I> \<Turnstile>\<^sub>c \<langle>S@map Send T@S',\<theta>\<rangle>\<close>
by (auto simp add: constr_sem_c_def)
next
case (Unify S f U \<delta> T S' \<theta>)
have "(\<theta> \<circ>\<^sub>s \<delta>) supports \<I>" "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
using Unify.prems(2) unfolding constr_sem_c_def by metis+
then obtain \<sigma> where \<sigma>: "\<theta> \<circ>\<^sub>s \<delta> \<circ>\<^sub>s \<sigma> = \<I>" unfolding subst_compose_def by auto
have \<theta>fun_id: "Fun f U \<cdot> \<theta> = Fun f U" "Fun f T \<cdot> \<theta> = Fun f T"
using Unify.prems(1) trm_subst_ident[of "Fun f U" \<theta>]
fv_subset_if_in_strand_ik[of "Fun f U" S] Unify.hyps(2)
fv_snd_rcv_strand_subset(2)[of S]
strand_vars_split(1)[of S "Send (Fun f T)#S'"]
unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def apply blast
using Unify.prems(1) trm_subst_ident[of "Fun f T" \<theta>]
unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce
hence \<theta>\<delta>_disj:
"subst_domain \<theta> \<inter> subst_domain \<delta> = {}"
"subst_domain \<theta> \<inter> range_vars \<delta> = {}"
"subst_domain \<theta> \<inter> range_vars \<theta> = {}"
using trm_subst_disj mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] apply (blast,blast)
using Unify.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast
hence \<theta>\<delta>_support: "\<theta> supports \<I>" "\<delta> supports \<I>"
by (simp_all add: subst_support_comp_split[OF \<open>(\<theta> \<circ>\<^sub>s \<delta>) supports \<I>\<close>])
have "fv (Fun f T) \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" "fv (Fun f U) \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using Unify.hyps(2) by force+
hence \<delta>_vars_bound: "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')"
using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by blast
have "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; [Send (Fun f T)]\<rbrakk>\<^sub>c \<I>"
proof -
from Unify.hyps(2) have "Fun f U \<cdot> \<I> \<in> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>" by blast
hence "Fun f U \<cdot> \<I> \<in> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>" by blast
moreover have "Unifier \<delta> (Fun f T) (Fun f U)"
by (fact MGU_is_Unifier[OF mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]])
ultimately have "Fun f T \<cdot> \<I> \<in> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>"
using \<sigma> by (metis \<theta>fun_id subst_subst_compose)
thus ?thesis by simp
qed
have "\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>"
proof -
have "(S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>s\<^sub>t \<theta> = S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "(S@S') \<cdot>\<^sub>s\<^sub>t \<theta> = S@S'"
proof -
have "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S') = {}"
using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
hence "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
using \<theta>\<delta>_disj(2) strand_subst_vars_union_bound[of "S@S'" \<delta>] by blast
thus "(S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>s\<^sub>t \<theta> = S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "(S@S') \<cdot>\<^sub>s\<^sub>t \<theta> = S@S'"
using strand_subst_comp \<open>subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S') = {}\<close> by (blast,blast)
qed
moreover have "subst_idem \<delta>" by (fact mgu_gives_subst_idem[OF Unify.hyps(3)[symmetric]])
moreover have
"(subst_domain \<theta> \<union> range_vars \<theta>) \<inter> bvars\<^sub>s\<^sub>t (S@S') = {}"
"(subst_domain \<theta> \<union> range_vars \<theta>) \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
"(subst_domain \<delta> \<union> range_vars \<delta>) \<inter> bvars\<^sub>s\<^sub>t (S@S') = {}"
using wf_constr_bvars_disj[OF Unify.prems(1)]
wf_constr_bvars_disj'[OF Unify.prems(1) \<delta>_vars_bound]
by auto
ultimately have "\<lbrakk>{}; S@S'\<rbrakk>\<^sub>c \<I>"
using \<open>\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>\<close> \<sigma>
strand_sem_subst(1)[of \<theta> "S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "{}" "\<delta> \<circ>\<^sub>s \<sigma>"]
strand_sem_subst(2)[of \<theta> "S@S'" "{}" "\<delta> \<circ>\<^sub>s \<sigma>"]
strand_sem_subst_subst_idem[of \<delta> "S@S'" "{}" \<sigma>]
unfolding constr_sem_c_def
by (metis subst_compose_assoc)
thus "\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>" by auto
qed
show "\<I> \<Turnstile>\<^sub>c \<langle>S@Send (Fun f T)#S',\<theta>\<rangle>"
using \<theta>\<delta>_support(1) \<open>\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; [Send (Fun f T)]\<rbrakk>\<^sub>c \<I>\<close> \<open>\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>\<close> \<open>\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>\<close>
by (auto simp add: constr_sem_c_def)
next
case (Equality S \<delta> t t' a S' \<theta>)
have "(\<theta> \<circ>\<^sub>s \<delta>) supports \<I>" "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
using Equality.prems(2) unfolding constr_sem_c_def by metis+
then obtain \<sigma> where \<sigma>: "\<theta> \<circ>\<^sub>s \<delta> \<circ>\<^sub>s \<sigma> = \<I>" unfolding subst_compose_def by auto
have "fv t \<subseteq> vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \<subseteq> vars\<^sub>s\<^sub>t (S@Equality a t t'#S')"
by auto
moreover have "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}"
using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by auto
ultimately have \<theta>fun_id: "t \<cdot> \<theta> = t" "t' \<cdot> \<theta> = t'"
using trm_subst_ident[of t \<theta>] trm_subst_ident[of t' \<theta>]
by auto
hence \<theta>\<delta>_disj:
"subst_domain \<theta> \<inter> subst_domain \<delta> = {}"
"subst_domain \<theta> \<inter> range_vars \<delta> = {}"
"subst_domain \<theta> \<inter> range_vars \<theta> = {}"
using trm_subst_disj mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] apply (blast,blast)
using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast
hence \<theta>\<delta>_support: "\<theta> supports \<I>" "\<delta> supports \<I>"
by (simp_all add: subst_support_comp_split[OF \<open>(\<theta> \<circ>\<^sub>s \<delta>) supports \<I>\<close>])
have "fv t \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto
hence \<delta>_vars_bound: "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t (S@Equality a t t'#S')"
using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by blast
have "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; [Equality a t t']\<rbrakk>\<^sub>c \<I>"
proof -
have "t \<cdot> \<delta> = t' \<cdot> \<delta>"
using MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]]
by metis
hence "t \<cdot> (\<theta> \<circ>\<^sub>s \<delta>) = t' \<cdot> (\<theta> \<circ>\<^sub>s \<delta>)" by (metis \<theta>fun_id subst_subst_compose)
hence "t \<cdot> \<I> = t' \<cdot> \<I>" by (metis \<sigma> subst_subst_compose)
thus ?thesis by simp
qed
have "\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>"
proof -
have "(S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>s\<^sub>t \<theta> = S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "(S@S') \<cdot>\<^sub>s\<^sub>t \<theta> = S@S'"
proof -
have "subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S') = {}"
using Equality.prems(1)
by (fastforce simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def simp del: subst_range.simps)
hence "subst_domain \<theta> \<inter> fv\<^sub>s\<^sub>t (S@S') = {}" by blast
hence "subst_domain \<theta> \<inter> fv\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
using \<theta>\<delta>_disj(2) subst_sends_strand_fv_to_img[of "S@S'" \<delta>] by blast
thus "(S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) \<cdot>\<^sub>s\<^sub>t \<theta> = S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "(S@S') \<cdot>\<^sub>s\<^sub>t \<theta> = S@S'"
using strand_subst_comp \<open>subst_domain \<theta> \<inter> vars\<^sub>s\<^sub>t (S@S') = {}\<close> by (blast,blast)
qed
moreover have
"(subst_domain \<theta> \<union> range_vars \<theta>) \<inter> bvars\<^sub>s\<^sub>t (S@S') = {}"
"(subst_domain \<theta> \<union> range_vars \<theta>) \<inter> bvars\<^sub>s\<^sub>t (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>) = {}"
"(subst_domain \<delta> \<union> range_vars \<delta>) \<inter> bvars\<^sub>s\<^sub>t (S@S') = {}"
using wf_constr_bvars_disj[OF Equality.prems(1)]
wf_constr_bvars_disj'[OF Equality.prems(1) \<delta>_vars_bound]
by auto
ultimately have "\<lbrakk>{}; S@S'\<rbrakk>\<^sub>c \<I>"
using \<open>\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>\<close> \<sigma>
strand_sem_subst(1)[of \<theta> "S@S' \<cdot>\<^sub>s\<^sub>t \<delta>" "{}" "\<delta> \<circ>\<^sub>s \<sigma>"]
strand_sem_subst(2)[of \<theta> "S@S'" "{}" "\<delta> \<circ>\<^sub>s \<sigma>"]
strand_sem_subst_subst_idem[of \<delta> "S@S'" "{}" \<sigma>]
mgu_gives_subst_idem[OF Equality.hyps(2)[symmetric]]
unfolding constr_sem_c_def
by (metis subst_compose_assoc)
thus "\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>" "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>" by auto
qed
show "\<I> \<Turnstile>\<^sub>c \<langle>S@Equality a t t'#S',\<theta>\<rangle>"
using \<theta>\<delta>_support(1) \<open>\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; [Equality a t t']\<rbrakk>\<^sub>c \<I>\<close> \<open>\<lbrakk>{}; S\<rbrakk>\<^sub>c \<I>\<close> \<open>\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>\<close>
by (auto simp add: constr_sem_c_def)
qed
theorem LI_soundness:
assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1" "(S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>2,\<theta>\<^sub>2)" "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>"
shows "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>"
using assms(2,1,3)
proof (induction S\<^sub>2 \<theta>\<^sub>2 rule: rtrancl_induct2)
case (step S\<^sub>i \<theta>\<^sub>i S\<^sub>j \<theta>\<^sub>j) thus ?case
using LI_preserves_wellformedness[OF \<open>(S\<^sub>1, \<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>i, \<theta>\<^sub>i)\<close> \<open>wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1\<close>]
LI_soundness_single[OF _ \<open>(S\<^sub>i, \<theta>\<^sub>i) \<leadsto> (S\<^sub>j, \<theta>\<^sub>j)\<close> \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>j, \<theta>\<^sub>j\<rangle>\<close>]
step.IH[OF \<open>wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1\<close>]
by metis
qed metis
end
subsection \<open>Theorem: Completeness of the Lazy Intruder\<close>
context
begin
private lemma LI_completeness_single:
assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1" "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>" "\<not>simple S\<^sub>1"
shows "\<exists>S\<^sub>2 \<theta>\<^sub>2. (S\<^sub>1,\<theta>\<^sub>1) \<leadsto> (S\<^sub>2,\<theta>\<^sub>2) \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>)"
using not_simple_elim[OF \<open>\<not>simple S\<^sub>1\<close>]
proof -
{ \<comment> \<open>In this case \<open>S\<^sub>1\<close> isn't simple because it contains an equality constraint,
so we can simply proceed with the reduction by computing the MGU for the equation\<close>
assume "\<exists>S' S'' a t t'. S\<^sub>1 = S'@Equality a t t'#S'' \<and> simple S'"
then obtain S a t t' S' where S\<^sub>1: "S\<^sub>1 = S@Equality a t t'#S'" "simple S" by moura
hence *: "wf\<^sub>s\<^sub>t {} S" "\<I> \<Turnstile>\<^sub>c \<langle>S, \<theta>\<^sub>1\<rangle>" "\<theta>\<^sub>1 supports \<I>" "t \<cdot> \<I> = t' \<cdot> \<I>"
using \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> \<open>wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1\<close> wf_eq_fv[of "{}" S t t' S']
fv_snd_rcv_strand_subset(5)[of S]
by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
from * have "Unifier \<I> t t'" by simp
then obtain \<delta> where \<delta>:
"Some \<delta> = mgu t t'" "subst_idem \<delta>" "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv t \<union> fv t'"
using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+
have "\<delta> \<preceq>\<^sub>\<circ> \<I>"
using mgu_gives_MGU[OF \<delta>(1)[symmetric]]
by (metis \<open>Unifier \<I> t t'\<close>)
hence "\<delta> supports \<I>" using subst_support_if_mgt_subst_idem[OF _ \<delta>(2)] by metis
hence "(\<theta>\<^sub>1 \<circ>\<^sub>s \<delta>) supports \<I>" using subst_support_comp \<open>\<theta>\<^sub>1 supports \<I>\<close> by metis
have "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
proof -
have "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t S\<^sub>1" using \<delta>(3) S\<^sub>1(1) by auto
hence "\<lbrakk>{}; S\<^sub>1 \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
using \<open>subst_idem \<delta>\<close> \<open>\<delta> \<preceq>\<^sub>\<circ> \<I>\<close> \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> strand_sem_subst
wf_constr_bvars_disj'(1)[OF assms(1)]
unfolding subst_idem_def constr_sem_c_def
by (metis (no_types) subst_compose_assoc)
thus "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>" using S\<^sub>1(1) by force
qed
moreover have "(S@Equality a t t'#S', \<theta>\<^sub>1) \<leadsto> (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>, \<theta>\<^sub>1 \<circ>\<^sub>s \<delta>)"
using LI_rel.Equality[OF \<open>simple S\<close> \<delta>(1)] S\<^sub>1 by metis
ultimately have ?thesis
using S\<^sub>1(1) \<open>(\<theta>\<^sub>1 \<circ>\<^sub>s \<delta>) supports \<I>\<close>
by (auto simp add: constr_sem_c_def)
} moreover {
\<comment> \<open>In this case \<open>S\<^sub>1\<close> isn't simple because it contains a deduction constraint for a composed
term, so we must look at how this composed term is derived under the interpretation \<open>\<I>\<close>\<close>
assume "\<exists>S' S'' f T. S\<^sub>1 = S'@Send (Fun f T)#S'' \<and> simple S'"
with assms obtain S f T S' where S\<^sub>1: "S\<^sub>1 = S@Send (Fun f T)#S'" "simple S" by moura
hence "wf\<^sub>s\<^sub>t {} S" "\<I> \<Turnstile>\<^sub>c \<langle>S, \<theta>\<^sub>1\<rangle>" "\<theta>\<^sub>1 supports \<I>"
using \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> \<open>wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1\<close>
by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def)
\<comment> \<open>Lemma for a common subcase\<close>
have fun_sat: "\<I> \<Turnstile>\<^sub>c \<langle>S@(map Send T)@S', \<theta>\<^sub>1\<rangle>" when T: "\<And>t. t \<in> set T \<Longrightarrow> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c t \<cdot> \<I>"
proof -
have "\<And>t. t \<in> set T \<Longrightarrow> \<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; [Send t]\<rbrakk>\<^sub>c \<I>" using T by simp
hence "\<lbrakk>ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; map Send T\<rbrakk>\<^sub>c \<I>" using \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> strand_sem_Send_map by metis
moreover have "\<lbrakk>ik\<^sub>s\<^sub>t (S@(map Send T)) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>; S'\<rbrakk>\<^sub>c \<I>"
using \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> S\<^sub>1
by (auto simp add: constr_sem_c_def)
ultimately show ?thesis
using \<open>\<I> \<Turnstile>\<^sub>c \<langle>S, \<theta>\<^sub>1\<rangle>\<close> \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close>
by (force simp add: constr_sem_c_def)
qed
from S\<^sub>1 \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> have "ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c Fun f T \<cdot> \<I>" by (auto simp add: constr_sem_c_def)
hence ?thesis
proof cases
\<comment> \<open>Case 1: \<open>\<I>(f(T))\<close> has been derived using the \<open>AxiomC\<close> rule.\<close>
case AxiomC
hence ex_t: "\<exists>t. t \<in> ik\<^sub>s\<^sub>t S \<and> Fun f T \<cdot> \<I> = t \<cdot> \<I>" by auto
show ?thesis
proof (cases "\<forall>T'. Fun f T' \<in> ik\<^sub>s\<^sub>t S \<longrightarrow> Fun f T \<cdot> \<I> \<noteq> Fun f T' \<cdot> \<I>")
\<comment> \<open>Case 1.1: \<open>f(T)\<close> is equal to a variable in the intruder knowledge under \<open>\<I>\<close>.
Hence there must exists a deduction constraint in the simple prefix of the constraint
in which this variable occurs/"is sent" for the first time. Since this variable itself
cannot have been derived from the \<open>AxiomC\<close> rule (because it must be equal under the
interpretation to \<open>f(T)\<close>, which is by assumption not in the intruder knowledge under
\<open>\<I>\<close>) it must be the case that we can derive it using the \<open>ComposeC\<close> rule. Hence we can
apply the \<open>Compose\<close> rule of the lazy intruder to \<open>f(T)\<close>.\<close>
case True
have "\<exists>v. Var v \<in> ik\<^sub>s\<^sub>t S \<and> Fun f T \<cdot> \<I> = \<I> v"
proof -
obtain t where "t \<in> ik\<^sub>s\<^sub>t S" "Fun f T \<cdot> \<I> = t \<cdot> \<I>" using ex_t by moura
thus ?thesis
using \<open>\<forall>T'. Fun f T' \<in> ik\<^sub>s\<^sub>t S \<longrightarrow> Fun f T \<cdot> \<I> \<noteq> Fun f T' \<cdot> \<I>\<close>
by (cases t) auto
qed
hence "\<exists>v \<in> wfrestrictedvars\<^sub>s\<^sub>t S. Fun f T \<cdot> \<I> = \<I> v"
using vars_subset_if_in_strand_ik2[of _ S] by fastforce
then obtain v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f
where S: "S = S\<^sub>p\<^sub>r\<^sub>e@Send (Var v)#S\<^sub>s\<^sub>u\<^sub>f" "Fun f T \<cdot> \<I> = \<I> v"
"\<not>(\<exists>w \<in> wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. Fun f T \<cdot> \<I> = \<I> w)"
using \<open>wf\<^sub>s\<^sub>t {} S\<close> wf_simple_strand_first_Send_var_split[OF _ \<open>simple S\<close>, of "Fun f T" \<I>]
by auto
hence "\<forall>w. Var w \<in> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<longrightarrow> \<I> v \<noteq> Var w \<cdot> \<I>" by auto
moreover have "\<forall>T'. Fun f T' \<in> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<longrightarrow> Fun f T \<cdot> \<I> \<noteq> Fun f T' \<cdot> \<I>"
using \<open>\<forall>T'. Fun f T' \<in> ik\<^sub>s\<^sub>t S \<longrightarrow> Fun f T \<cdot> \<I> \<noteq> Fun f T' \<cdot> \<I>\<close> S(1)
by (meson contra_subsetD ik_append_subset(1))
hence "\<forall>g T'. Fun g T' \<in> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<longrightarrow> \<I> v \<noteq> Fun g T' \<cdot> \<I>" using S(2) by simp
ultimately have "\<forall>t \<in> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. \<I> v \<noteq> t \<cdot> \<I>" by (metis term.exhaust)
hence "\<I> v \<notin> (ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>" by auto
have "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c \<I> v"
using S\<^sub>1(1) S(1) \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close>
by (auto simp add: constr_sem_c_def)
hence "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c Fun f T \<cdot> \<I>" using \<open>Fun f T \<cdot> \<I> = \<I> v\<close> by metis
hence "length T = arity f" "public f" "\<And>t. t \<in> set T \<Longrightarrow> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c t \<cdot> \<I>"
using \<open>Fun f T \<cdot> \<I> = \<I> v\<close> \<open>\<I> v \<notin> ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>\<close>
intruder_synth.simps[of "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I>" "\<I> v"]
by auto
hence *: "\<And>t. t \<in> set T \<Longrightarrow> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c t \<cdot> \<I>"
using S(1) by (auto intro: ideduct_synth_mono)
hence "\<I> \<Turnstile>\<^sub>c \<langle>S@(map Send T)@S', \<theta>\<^sub>1\<rangle>" by (metis fun_sat)
moreover have "(S@Send (Fun f T)#S', \<theta>\<^sub>1) \<leadsto> (S@map Send T@S', \<theta>\<^sub>1)"
by (metis LI_rel.Compose[OF \<open>simple S\<close> \<open>length T = arity f\<close> \<open>public f\<close>])
ultimately show ?thesis using S\<^sub>1 by auto
next
\<comment> \<open>Case 1.2: \<open>\<I>(f(T))\<close> can be derived from an interpreted composed term in the intruder
knowledge. Use the \<open>Unify\<close> rule on this composed term to further reduce the constraint.\<close>
case False
then obtain T' where t: "Fun f T' \<in> ik\<^sub>s\<^sub>t S" "Fun f T \<cdot> \<I> = Fun f T' \<cdot> \<I>"
by auto
hence "fv (Fun f T') \<subseteq> fv\<^sub>s\<^sub>t S\<^sub>1"
using S\<^sub>1(1) fv_subset_if_in_strand_ik'[OF t(1)]
fv_snd_rcv_strand_subset(2)[of S]
by auto
from t have "Unifier \<I> (Fun f T) (Fun f T')" by simp
then obtain \<delta> where \<delta>:
"Some \<delta> = mgu (Fun f T) (Fun f T')" "subst_idem \<delta>"
"subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv (Fun f T) \<union> fv (Fun f T')"
using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+
have "\<delta> \<preceq>\<^sub>\<circ> \<I>"
using mgu_gives_MGU[OF \<delta>(1)[symmetric]]
by (metis \<open>Unifier \<I> (Fun f T) (Fun f T')\<close>)
hence "\<delta> supports \<I>" using subst_support_if_mgt_subst_idem[OF _ \<delta>(2)] by metis
hence "(\<theta>\<^sub>1 \<circ>\<^sub>s \<delta>) supports \<I>" using subst_support_comp \<open>\<theta>\<^sub>1 supports \<I>\<close> by metis
have "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
proof -
have "subst_domain \<delta> \<union> range_vars \<delta> \<subseteq> fv\<^sub>s\<^sub>t S\<^sub>1"
using \<delta>(3) S\<^sub>1(1) \<open>fv (Fun f T') \<subseteq> fv\<^sub>s\<^sub>t S\<^sub>1\<close>
unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps)
hence "\<lbrakk>{}; S\<^sub>1 \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>"
using \<open>subst_idem \<delta>\<close> \<open>\<delta> \<preceq>\<^sub>\<circ> \<I>\<close> \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close> strand_sem_subst
wf_constr_bvars_disj'(1)[OF assms(1)]
unfolding subst_idem_def constr_sem_c_def
by (metis (no_types) subst_compose_assoc)
thus "\<lbrakk>{}; S@S' \<cdot>\<^sub>s\<^sub>t \<delta>\<rbrakk>\<^sub>c \<I>" using S\<^sub>1(1) by force
qed
moreover have "(S@Send (Fun f T)#S', \<theta>\<^sub>1) \<leadsto> (S@S' \<cdot>\<^sub>s\<^sub>t \<delta>, \<theta>\<^sub>1 \<circ>\<^sub>s \<delta>)"
using LI_rel.Unify[OF \<open>simple S\<close> t(1) \<delta>(1)] S\<^sub>1 by metis
ultimately show ?thesis
using S\<^sub>1(1) \<open>(\<theta>\<^sub>1 \<circ>\<^sub>s \<delta>) supports \<I>\<close>
by (auto simp add: constr_sem_c_def)
qed
next
\<comment> \<open>Case 2: \<open>\<I>(f(T))\<close> has been derived using the \<open>ComposeC\<close> rule.
Simply use the \<open>Compose\<close> rule of the lazy intruder to proceed with the reduction.\<close>
case (ComposeC T' g)
hence "f = g" "length T = arity f" "public f"
and "\<And>x. x \<in> set T \<Longrightarrow> ik\<^sub>s\<^sub>t S \<cdot>\<^sub>s\<^sub>e\<^sub>t \<I> \<turnstile>\<^sub>c x \<cdot> \<I>"
by auto
hence "\<I> \<Turnstile>\<^sub>c \<langle>S@(map Send T)@S', \<theta>\<^sub>1\<rangle>" using fun_sat by metis
moreover have "(S\<^sub>1, \<theta>\<^sub>1) \<leadsto> (S@(map Send T)@S', \<theta>\<^sub>1)"
using S\<^sub>1 LI_rel.Compose[OF \<open>simple S\<close> \<open>length T = arity f\<close> \<open>public f\<close>]
by metis
ultimately show ?thesis by metis
qed
} moreover have "\<And>A B X F. S\<^sub>1 = A@Inequality X F#B \<Longrightarrow> ineq_model \<I> X F"
using assms(2) by (auto simp add: constr_sem_c_def)
ultimately show ?thesis using not_simple_elim[OF \<open>\<not>simple S\<^sub>1\<close>] by metis
qed
theorem LI_completeness:
assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1" "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>"
shows "\<exists>S\<^sub>2 \<theta>\<^sub>2. (S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>2,\<theta>\<^sub>2) \<and> simple S\<^sub>2 \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>)"
proof (cases "simple S\<^sub>1")
case False
let ?Stuck = "\<lambda>S\<^sub>2 \<theta>\<^sub>2. \<not>(\<exists>S\<^sub>3 \<theta>\<^sub>3. (S\<^sub>2,\<theta>\<^sub>2) \<leadsto> (S\<^sub>3,\<theta>\<^sub>3) \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>3, \<theta>\<^sub>3\<rangle>))"
let ?Sats = "{((S,\<theta>),(S',\<theta>')). (S,\<theta>) \<leadsto> (S',\<theta>') \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S, \<theta>\<rangle>) \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S', \<theta>'\<rangle>)}"
have simple_if_stuck:
"\<And>S\<^sub>2 \<theta>\<^sub>2. \<lbrakk>(S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>+ (S\<^sub>2,\<theta>\<^sub>2); \<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>; ?Stuck S\<^sub>2 \<theta>\<^sub>2\<rbrakk> \<Longrightarrow> simple S\<^sub>2"
using LI_completeness_single
LI_preserves_wellformedness
\<open>wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1\<close>
trancl_into_rtrancl
by metis
have base: "\<exists>b. ((S\<^sub>1,\<theta>\<^sub>1),b) \<in> ?Sats"
using LI_completeness_single[OF assms False] assms(2)
by auto
have *: "\<And>S \<theta> S' \<theta>'. ((S,\<theta>),(S',\<theta>')) \<in> ?Sats\<^sup>+ \<Longrightarrow> (S,\<theta>) \<leadsto>\<^sup>+ (S',\<theta>') \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S', \<theta>'\<rangle>)"
proof -
fix S \<theta> S' \<theta>'
assume "((S,\<theta>),(S',\<theta>')) \<in> ?Sats\<^sup>+"
thus "(S,\<theta>) \<leadsto>\<^sup>+ (S',\<theta>') \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S', \<theta>'\<rangle>)"
by (induct rule: trancl_induct2) auto
qed
have "\<exists>S\<^sub>2 \<theta>\<^sub>2. ((S\<^sub>1,\<theta>\<^sub>1),(S\<^sub>2,\<theta>\<^sub>2)) \<in> ?Sats\<^sup>+ \<and> ?Stuck S\<^sub>2 \<theta>\<^sub>2"
proof (rule ccontr)
assume "\<not>(\<exists>S\<^sub>2 \<theta>\<^sub>2. ((S\<^sub>1,\<theta>\<^sub>1),(S\<^sub>2,\<theta>\<^sub>2)) \<in> ?Sats\<^sup>+ \<and> ?Stuck S\<^sub>2 \<theta>\<^sub>2)"
hence sat_not_stuck: "\<And>S\<^sub>2 \<theta>\<^sub>2. ((S\<^sub>1,\<theta>\<^sub>1),(S\<^sub>2,\<theta>\<^sub>2)) \<in> ?Sats\<^sup>+ \<Longrightarrow> \<not>?Stuck S\<^sub>2 \<theta>\<^sub>2" by blast
have "\<forall>S \<theta>. ((S\<^sub>1,\<theta>\<^sub>1),(S,\<theta>)) \<in> ?Sats\<^sup>+ \<longrightarrow> (\<exists>b. ((S,\<theta>),b) \<in> ?Sats)"
proof (intro allI impI)
fix S \<theta> assume a: "((S\<^sub>1,\<theta>\<^sub>1),(S,\<theta>)) \<in> ?Sats\<^sup>+"
have "\<And>b. ((S\<^sub>1,\<theta>\<^sub>1),b) \<in> ?Sats\<^sup>+ \<Longrightarrow> \<exists>c. b \<leadsto> c \<and> ((S\<^sub>1,\<theta>\<^sub>1),c) \<in> ?Sats\<^sup>+"
proof -
fix b assume in_sat: "((S\<^sub>1,\<theta>\<^sub>1),b) \<in> ?Sats\<^sup>+"
hence "\<exists>c. (b,c) \<in> ?Sats" using * sat_not_stuck by (cases b) blast
thus "\<exists>c. b \<leadsto> c \<and> ((S\<^sub>1,\<theta>\<^sub>1),c) \<in> ?Sats\<^sup>+"
using trancl_into_trancl[OF in_sat] by blast
qed
hence "\<exists>S' \<theta>'. (S,\<theta>) \<leadsto> (S',\<theta>') \<and> ((S\<^sub>1,\<theta>\<^sub>1),(S',\<theta>')) \<in> ?Sats\<^sup>+" using a by auto
then obtain S' \<theta>' where S'\<theta>': "(S,\<theta>) \<leadsto> (S',\<theta>')" "((S\<^sub>1,\<theta>\<^sub>1),(S',\<theta>')) \<in> ?Sats\<^sup>+" by auto
hence "\<I> \<Turnstile>\<^sub>c \<langle>S', \<theta>'\<rangle>" using * by blast
moreover have "(S\<^sub>1, \<theta>\<^sub>1) \<leadsto>\<^sup>+ (S,\<theta>)" using a trancl_mono by blast
ultimately have "((S,\<theta>),(S',\<theta>')) \<in> ?Sats" using S'\<theta>'(1) * a by blast
thus "\<exists>b. ((S,\<theta>),b) \<in> ?Sats" using S'\<theta>'(2) by blast
qed
hence "\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> ?Sats"
using infinite_chain_intro'[OF base] by blast
moreover have "?Sats \<subseteq> LI_rel\<^sup>+" by auto
hence "\<not>(\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> ?Sats)"
using LI_no_infinite_chain infinite_chain_mono by blast
ultimately show False by auto
qed
hence "\<exists>S\<^sub>2 \<theta>\<^sub>2. (S\<^sub>1, \<theta>\<^sub>1) \<leadsto>\<^sup>+ (S\<^sub>2, \<theta>\<^sub>2) \<and> simple S\<^sub>2 \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>)"
using simple_if_stuck * by blast
thus ?thesis by (meson trancl_into_rtrancl)
qed (blast intro: \<open>\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle>\<close>)
end
subsection \<open>Corollary: Soundness and Completeness as a Single Theorem\<close>
corollary LI_soundness_and_completeness:
assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \<theta>\<^sub>1"
shows "\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>1, \<theta>\<^sub>1\<rangle> \<longleftrightarrow> (\<exists>S\<^sub>2 \<theta>\<^sub>2. (S\<^sub>1,\<theta>\<^sub>1) \<leadsto>\<^sup>* (S\<^sub>2,\<theta>\<^sub>2) \<and> simple S\<^sub>2 \<and> (\<I> \<Turnstile>\<^sub>c \<langle>S\<^sub>2, \<theta>\<^sub>2\<rangle>))"
by (metis LI_soundness[OF assms] LI_completeness[OF assms])
end
end

View File

@ -0,0 +1,538 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Messages.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>Protocol Messages as (First-Order) Terms\<close>
theory Messages
imports Miscellaneous "First_Order_Terms.Term"
begin
subsection \<open>Term-related definitions: subterms and free variables\<close>
abbreviation "the_Fun \<equiv> un_Fun1"
lemmas the_Fun_def = un_Fun1_def
fun subterms::"('a,'b) term \<Rightarrow> ('a,'b) terms" where
"subterms (Var x) = {Var x}"
| "subterms (Fun f T) = {Fun f T} \<union> (\<Union>t \<in> set T. subterms t)"
abbreviation subtermeq (infix "\<sqsubseteq>" 50) where "t' \<sqsubseteq> t \<equiv> (t' \<in> subterms t)"
abbreviation subterm (infix "\<sqsubset>" 50) where "t' \<sqsubset> t \<equiv> (t' \<sqsubseteq> t \<and> t' \<noteq> t)"
abbreviation "subterms\<^sub>s\<^sub>e\<^sub>t M \<equiv> \<Union>(subterms ` M)"
abbreviation subtermeqset (infix "\<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t" 50) where "t \<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t M \<equiv> (t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M)"
abbreviation fv where "fv \<equiv> vars_term"
lemmas fv_simps = term.simps(17,18)
fun fv\<^sub>s\<^sub>e\<^sub>t where "fv\<^sub>s\<^sub>e\<^sub>t M = \<Union>(fv ` M)"
abbreviation fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r where "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r p \<equiv> case p of (t,t') \<Rightarrow> fv t \<union> fv t'"
fun fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s where "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = \<Union>(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set F)"
abbreviation ground where "ground M \<equiv> fv\<^sub>s\<^sub>e\<^sub>t M = {}"
subsection \<open>Variants that return lists insteads of sets\<close>
fun fv_list where
"fv_list (Var x) = [x]"
| "fv_list (Fun f T) = concat (map fv_list T)"
definition fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s where
"fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<equiv> concat (map (\<lambda>(t,t'). fv_list t@fv_list t') F)"
fun subterms_list::"('a,'b) term \<Rightarrow> ('a,'b) term list" where
"subterms_list (Var x) = [Var x]"
| "subterms_list (Fun f T) = remdups (Fun f T#concat (map subterms_list T))"
lemma fv_list_is_fv: "fv t = set (fv_list t)"
by (induct t) auto
lemma fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = set (fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)"
by (induct F) (auto simp add: fv_list_is_fv fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_def)
lemma subterms_list_is_subterms: "subterms t = set (subterms_list t)"
by (induct t) auto
subsection \<open>The subterm relation defined as a function\<close>
fun subterm_of where
"subterm_of t (Var y) = (t = Var y)"
| "subterm_of t (Fun f T) = (t = Fun f T \<or> list_ex (subterm_of t) T)"
lemma subterm_of_iff_subtermeq[code_unfold]: "t \<sqsubseteq> t' = subterm_of t t'"
proof (induction t')
case (Fun f T) thus ?case
proof (cases "t = Fun f T")
case False thus ?thesis
using Fun.IH subterm_of.simps(2)[of t f T]
unfolding list_ex_iff by fastforce
qed simp
qed simp
lemma subterm_of_ex_set_iff_subtermeqset[code_unfold]: "t \<sqsubseteq>\<^sub>s\<^sub>e\<^sub>t M = (\<exists>t' \<in> M. subterm_of t t')"
using subterm_of_iff_subtermeq by blast
subsection \<open>The subterm relation is a partial order on terms\<close>
interpretation "term": order "(\<sqsubseteq>)" "(\<sqsubset>)"
proof
show "s \<sqsubseteq> s" for s :: "('a,'b) term"
by (induct s rule: subterms.induct) auto
show trans: "s \<sqsubseteq> t \<Longrightarrow> t \<sqsubseteq> u \<Longrightarrow> s \<sqsubseteq> u" for s t u :: "('a,'b) term"
by (induct u rule: subterms.induct) auto
show "s \<sqsubseteq> t \<Longrightarrow> t \<sqsubseteq> s \<Longrightarrow> s = t" for s t :: "('a,'b) term"
proof (induction s arbitrary: t rule: subterms.induct[case_names Var Fun])
case (Fun f T)
{ assume 0: "t \<noteq> Fun f T"
then obtain u::"('a,'b) term" where u: "u \<in> set T" "t \<sqsubseteq> u" using Fun.prems(2) by auto
hence 1: "Fun f T \<sqsubseteq> u" using trans[OF Fun.prems(1)] by simp
have 2: "u \<sqsubseteq> Fun f T"
by (cases u) (use u(1) in force, use u(1) subterms.simps(2)[of f T] in fastforce)
hence 3: "u = Fun f T" using Fun.IH[OF u(1) _ 1] by simp
have "u \<sqsubseteq> t" using trans[OF 2 Fun.prems(1)] by simp
hence 4: "u = t" using Fun.IH[OF u(1) _ u(2)] by simp
have "t = Fun f T" using 3 4 by simp
hence False using 0 by simp
}
thus ?case by auto
qed simp
thus "(s \<sqsubset> t) = (s \<sqsubseteq> t \<and> \<not>(t \<sqsubseteq> s))" for s t :: "('a,'b) term"
by blast
qed
subsection \<open>Lemmata concerning subterms and free variables\<close>
lemma fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append: "fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) = fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F@fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G"
by (simp add: fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_def)
lemma distinct_fv_list_idx_fv_disjoint:
assumes t: "distinct (fv_list t)" "Fun f T \<sqsubseteq> t"
and ij: "i < length T" "j < length T" "i < j"
shows "fv (T ! i) \<inter> fv (T ! j) = {}"
using t
proof (induction t rule: fv_list.induct)
case (2 g S)
have "distinct (fv_list s)" when s: "s \<in> set S" for s
by (metis (no_types, lifting) s "2.prems"(1) concat_append distinct_append
map_append split_list fv_list.simps(2) concat.simps(2) list.simps(9))
hence IH: "fv (T ! i) \<inter> fv (T ! j) = {}"
when s: "s \<in> set S" "Fun f T \<sqsubseteq> s" for s
using "2.IH" s by blast
show ?case
proof (cases "Fun f T = Fun g S")
case True
define U where "U \<equiv> map fv_list T"
have a: "distinct (concat U)"
using "2.prems"(1) True unfolding U_def by auto
have b: "i < length U" "j < length U"
using ij(1,2) unfolding U_def by simp_all
show ?thesis
using b distinct_concat_idx_disjoint[OF a b ij(3)]
fv_list_is_fv[of "T ! i"] fv_list_is_fv[of "T ! j"]
unfolding U_def by force
qed (use IH "2.prems"(2) in auto)
qed force
lemmas subtermeqI'[intro] = term.eq_refl
lemma subtermeqI''[intro]: "t \<in> set T \<Longrightarrow> t \<sqsubseteq> Fun f T"
by force
lemma finite_fv_set[intro]: "finite M \<Longrightarrow> finite (fv\<^sub>s\<^sub>e\<^sub>t M)"
by auto
lemma finite_fun_symbols[simp]: "finite (funs_term t)"
by (induct t) simp_all
lemma fv_set_mono: "M \<subseteq> N \<Longrightarrow> fv\<^sub>s\<^sub>e\<^sub>t M \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma subterms\<^sub>s\<^sub>e\<^sub>t_mono: "M \<subseteq> N \<Longrightarrow> subterms\<^sub>s\<^sub>e\<^sub>t M \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma ground_empty[simp]: "ground {}"
by simp
lemma ground_subset: "M \<subseteq> N \<Longrightarrow> ground N \<Longrightarrow> ground M"
by auto
lemma fv_map_fv_set: "\<Union>(set (map fv L)) = fv\<^sub>s\<^sub>e\<^sub>t (set L)"
by (induct L) auto
lemma fv\<^sub>s\<^sub>e\<^sub>t_union: "fv\<^sub>s\<^sub>e\<^sub>t (M \<union> N) = fv\<^sub>s\<^sub>e\<^sub>t M \<union> fv\<^sub>s\<^sub>e\<^sub>t N"
by auto
lemma finite_subset_Union:
fixes A::"'a set" and f::"'a \<Rightarrow> 'b set"
assumes "finite (\<Union>a \<in> A. f a)"
shows "\<exists>B. finite B \<and> B \<subseteq> A \<and> (\<Union>b \<in> B. f b) = (\<Union>a \<in> A. f a)"
by (metis assms eq_iff finite_subset_image finite_UnionD)
lemma inv_set_fv: "finite M \<Longrightarrow> \<Union>(set (map fv (inv set M))) = fv\<^sub>s\<^sub>e\<^sub>t M"
using fv_map_fv_set[of "inv set M"] inv_set_fset by auto
lemma ground_subterm: "fv t = {} \<Longrightarrow> t' \<sqsubseteq> t \<Longrightarrow> fv t' = {}" by (induct t) auto
lemma empty_fv_not_var: "fv t = {} \<Longrightarrow> t \<noteq> Var x" by auto
lemma empty_fv_exists_fun: "fv t = {} \<Longrightarrow> \<exists>f X. t = Fun f X" by (cases t) auto
lemma vars_iff_subtermeq: "x \<in> fv t \<longleftrightarrow> Var x \<sqsubseteq> t" by (induct t) auto
lemma vars_iff_subtermeq_set: "x \<in> fv\<^sub>s\<^sub>e\<^sub>t M \<longleftrightarrow> Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M"
using vars_iff_subtermeq[of x] by auto
lemma vars_if_subtermeq_set: "Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> x \<in> fv\<^sub>s\<^sub>e\<^sub>t M"
by (metis vars_iff_subtermeq_set)
lemma subtermeq_set_if_vars: "x \<in> fv\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> Var x \<in> subterms\<^sub>s\<^sub>e\<^sub>t M"
by (metis vars_iff_subtermeq_set)
lemma vars_iff_subterm_or_eq: "x \<in> fv t \<longleftrightarrow> Var x \<sqsubset> t \<or> Var x = t"
by (induct t) (auto simp add: vars_iff_subtermeq)
lemma var_is_subterm: "x \<in> fv t \<Longrightarrow> Var x \<in> subterms t"
by (simp add: vars_iff_subtermeq)
lemma subterm_is_var: "Var x \<in> subterms t \<Longrightarrow> x \<in> fv t"
by (simp add: vars_iff_subtermeq)
lemma no_var_subterm: "\<not>t \<sqsubset> Var v" by auto
lemma fun_if_subterm: "t \<sqsubset> u \<Longrightarrow> \<exists>f X. u = Fun f X" by (induct u) simp_all
lemma subtermeq_vars_subset: "M \<sqsubseteq> N \<Longrightarrow> fv M \<subseteq> fv N" by (induct N) auto
lemma fv_subterms[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms t) = fv t"
by (induct t) auto
lemma fv_subterms_set[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M) = fv\<^sub>s\<^sub>e\<^sub>t M"
using subtermeq_vars_subset by auto
lemma fv_subset: "t \<in> M \<Longrightarrow> fv t \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t M"
by auto
lemma fv_subset_subterms: "t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> fv t \<subseteq> fv\<^sub>s\<^sub>e\<^sub>t M"
using fv_subset fv_subterms_set by metis
lemma subterms_finite[simp]: "finite (subterms t)" by (induction rule: subterms.induct) auto
lemma subterms_union_finite: "finite M \<Longrightarrow> finite (\<Union>t \<in> M. subterms t)"
by (induction rule: subterms.induct) auto
lemma subterms_subset: "t' \<sqsubseteq> t \<Longrightarrow> subterms t' \<subseteq> subterms t"
by (induction rule: subterms.induct) auto
lemma subterms_subset_set: "M \<subseteq> subterms t \<Longrightarrow> subterms\<^sub>s\<^sub>e\<^sub>t M \<subseteq> subterms t"
by (metis SUP_least contra_subsetD subterms_subset)
lemma subset_subterms_Union[simp]: "M \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t M" by auto
lemma in_subterms_Union: "t \<in> M \<Longrightarrow> t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M" using subset_subterms_Union by blast
lemma in_subterms_subset_Union: "t \<in> subterms\<^sub>s\<^sub>e\<^sub>t M \<Longrightarrow> subterms t \<subseteq> subterms\<^sub>s\<^sub>e\<^sub>t M"
using subterms_subset by auto
lemma subterm_param_split:
assumes "t \<sqsubset> Fun f X"
shows "\<exists>pre x suf. t \<sqsubseteq> x \<and> X = pre@x#suf"
proof -
obtain x where "t \<sqsubseteq> x" "x \<in> set X" using assms by auto
then obtain pre suf where "X = pre@x#suf" "x \<notin> set pre \<or> x \<notin> set suf"
by (meson split_list_first split_list_last)
thus ?thesis using \<open>t \<sqsubseteq> x\<close> by auto
qed
lemma ground_iff_no_vars: "ground (M::('a,'b) terms) \<longleftrightarrow> (\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m))"
proof
assume "ground M"
hence "\<forall>v. \<forall>m \<in> M. v \<notin> fv m" by auto
hence "\<forall>v. \<forall>m \<in> M. Var v \<notin> subterms m" by (simp add: vars_iff_subtermeq)
thus "(\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m))" by simp
next
assume no_vars: "\<forall>v. Var v \<notin> (\<Union>m \<in> M. subterms m)"
moreover
{ assume "\<not>ground M"
then obtain v and m::"('a,'b) term" where "m \<in> M" "fv m \<noteq> {}" "v \<in> fv m" by auto
hence "Var v \<in> (subterms m)" by (simp add: vars_iff_subtermeq)
hence "\<exists>v. Var v \<in> (\<Union>t \<in> M. subterms t)" using \<open>m \<in> M\<close> by auto
hence False using no_vars by simp
}
ultimately show "ground M" by blast
qed
lemma index_Fun_subterms_subset[simp]: "i < length T \<Longrightarrow> subterms (T ! i) \<subseteq> subterms (Fun f T)"
by auto
lemma index_Fun_fv_subset[simp]: "i < length T \<Longrightarrow> fv (T ! i) \<subseteq> fv (Fun f T)"
using subtermeq_vars_subset by fastforce
lemma subterms_union_ground:
assumes "ground M"
shows "ground (subterms\<^sub>s\<^sub>e\<^sub>t M)"
proof -
{ fix t assume "t \<in> M"
hence "fv t = {}"
using ground_iff_no_vars[of M] assms
by auto
hence "\<forall>t' \<in> subterms t. fv t' = {}" using subtermeq_vars_subset[of _ t] by simp
hence "ground (subterms t)" by auto
}
thus ?thesis by auto
qed
lemma Var_subtermeq: "t \<sqsubseteq> Var v \<Longrightarrow> t = Var v" by simp
lemma subtermeq_imp_funs_term_subset: "s \<sqsubseteq> t \<Longrightarrow> funs_term s \<subseteq> funs_term t"
by (induct t arbitrary: s) auto
lemma subterms_const: "subterms (Fun f []) = {Fun f []}" by simp
lemma subterm_subtermeq_neq: "\<lbrakk>t \<sqsubset> u; u \<sqsubseteq> v\<rbrakk> \<Longrightarrow> t \<noteq> v"
by (metis term.eq_iff)
lemma subtermeq_subterm_neq: "\<lbrakk>t \<sqsubseteq> u; u \<sqsubset> v\<rbrakk> \<Longrightarrow> t \<noteq> v"
by (metis term.eq_iff)
lemma subterm_size_lt: "x \<sqsubset> y \<Longrightarrow> size x < size y"
using not_less_eq size_list_estimation by (induct y, simp, fastforce)
lemma in_subterms_eq: "\<lbrakk>x \<in> subterms y; y \<in> subterms x\<rbrakk> \<Longrightarrow> subterms x = subterms y"
using term.antisym by auto
lemma Fun_gt_params: "Fun f X \<notin> (\<Union>x \<in> set X. subterms x)"
proof -
have "size_list size X < size (Fun f X)" by simp
hence "Fun f X \<notin> set X" by (meson less_not_refl size_list_estimation)
hence "\<forall>x \<in> set X. Fun f X \<notin> subterms x \<or> x \<notin> subterms (Fun f X)"
by (metis term.antisym[of "Fun f X" _])
moreover have "\<forall>x \<in> set X. x \<in> subterms (Fun f X)" by fastforce
ultimately show ?thesis by auto
qed
lemma params_subterms[simp]: "set X \<subseteq> subterms (Fun f X)" by auto
lemma params_subterms_Union[simp]: "subterms\<^sub>s\<^sub>e\<^sub>t (set X) \<subseteq> subterms (Fun f X)" by auto
lemma Fun_subterm_inside_params: "t \<sqsubset> Fun f X \<longleftrightarrow> t \<in> (\<Union>x \<in> (set X). subterms x)"
using Fun_gt_params by fastforce
lemma Fun_param_is_subterm: "x \<in> set X \<Longrightarrow> x \<sqsubset> Fun f X"
using Fun_subterm_inside_params by fastforce
lemma Fun_param_in_subterms: "x \<in> set X \<Longrightarrow> x \<in> subterms (Fun f X)"
using Fun_subterm_inside_params by fastforce
lemma Fun_not_in_param: "x \<in> set X \<Longrightarrow> \<not>Fun f X \<sqsubset> x"
using term.antisym by fast
lemma Fun_ex_if_subterm: "t \<sqsubset> s \<Longrightarrow> \<exists>f T. Fun f T \<sqsubseteq> s \<and> t \<in> set T"
proof (induction s)
case (Fun f T)
then obtain s' where s': "s' \<in> set T" "t \<sqsubseteq> s'" by auto
show ?case
proof (cases "t = s'")
case True thus ?thesis using s' by blast
next
case False
thus ?thesis
using Fun.IH[OF s'(1)] s'(2) term.order_trans[OF _ Fun_param_in_subterms[OF s'(1), of f]]
by metis
qed
qed simp
lemma const_subterm_obtain:
assumes "fv t = {}"
obtains c where "Fun c [] \<sqsubseteq> t"
using assms
proof (induction t)
case (Fun f T) thus ?case by (cases "T = []") force+
qed simp
lemma const_subterm_obtain': "fv t = {} \<Longrightarrow> \<exists>c. Fun c [] \<sqsubseteq> t"
by (metis const_subterm_obtain)
lemma subterms_singleton:
assumes "(\<exists>v. t = Var v) \<or> (\<exists>f. t = Fun f [])"
shows "subterms t = {t}"
using assms by (cases t) auto
lemma subtermeq_Var_const:
assumes "s \<sqsubseteq> t"
shows "t = Var v \<Longrightarrow> s = Var v" "t = Fun f [] \<Longrightarrow> s = Fun f []"
using assms by fastforce+
lemma subterms_singleton':
assumes "subterms t = {t}"
shows "(\<exists>v. t = Var v) \<or> (\<exists>f. t = Fun f [])"
proof (cases t)
case (Fun f T)
{ fix s S assume "T = s#S"
hence "s \<in> subterms t" using Fun by auto
hence "s = t" using assms by auto
hence False
using Fun_param_is_subterm[of s "s#S" f] \<open>T = s#S\<close> Fun
by auto
}
hence "T = []" by (cases T) auto
thus ?thesis using Fun by simp
qed (simp add: assms)
lemma funs_term_subterms_eq[simp]:
"(\<Union>s \<in> subterms t. funs_term s) = funs_term t"
"(\<Union>s \<in> subterms\<^sub>s\<^sub>e\<^sub>t M. funs_term s) = \<Union>(funs_term ` M)"
proof -
show "\<And>t. \<Union>(funs_term ` subterms t) = funs_term t"
using term.order_refl subtermeq_imp_funs_term_subset by blast
thus "\<Union>(funs_term ` (subterms\<^sub>s\<^sub>e\<^sub>t M)) = \<Union>(funs_term ` M)" by force
qed
lemmas subtermI'[intro] = Fun_param_is_subterm
lemma funs_term_Fun_subterm: "f \<in> funs_term t \<Longrightarrow> \<exists>T. Fun f T \<in> subterms t"
proof (induction t)
case (Fun g T)
hence "f = g \<or> (\<exists>s \<in> set T. f \<in> funs_term s)" by simp
thus ?case
proof
assume "\<exists>s \<in> set T. f \<in> funs_term s"
then obtain s where "s \<in> set T" "\<exists>T. Fun f T \<in> subterms s" using Fun.IH by auto
thus ?thesis by auto
qed (auto simp add: Fun)
qed simp
lemma funs_term_Fun_subterm': "Fun f T \<in> subterms t \<Longrightarrow> f \<in> funs_term t"
by (induct t) auto
lemma zip_arg_subterm:
assumes "(s,t) \<in> set (zip X Y)"
shows "s \<sqsubset> Fun f X" "t \<sqsubset> Fun g Y"
proof -
from assms have *: "s \<in> set X" "t \<in> set Y" by (meson in_set_zipE)+
show "s \<sqsubset> Fun f X" by (metis Fun_param_is_subterm[OF *(1)])
show "t \<sqsubset> Fun g Y" by (metis Fun_param_is_subterm[OF *(2)])
qed
lemma fv_disj_Fun_subterm_param_cases:
assumes "fv t \<inter> X = {}" "Fun f T \<in> subterms t"
shows "T = [] \<or> (\<exists>s\<in>set T. s \<notin> Var ` X)"
proof (cases T)
case (Cons s S)
hence "s \<in> subterms t"
using assms(2) term.order_trans[of _ "Fun f T" t]
by auto
hence "fv s \<inter> X = {}" using assms(1) fv_subterms by force
thus ?thesis using Cons by auto
qed simp
lemma fv_eq_FunI:
assumes "length T = length S" "\<And>i. i < length T \<Longrightarrow> fv (T ! i) = fv (S ! i)"
shows "fv (Fun f T) = fv (Fun g S)"
using assms
proof (induction T arbitrary: S)
case (Cons t T S')
then obtain s S where S': "S' = s#S" by (cases S') simp_all
thus ?case using Cons by fastforce
qed simp
lemma fv_eq_FunI':
assumes "length T = length S" "\<And>i. i < length T \<Longrightarrow> x \<in> fv (T ! i) \<longleftrightarrow> x \<in> fv (S ! i)"
shows "x \<in> fv (Fun f T) \<longleftrightarrow> x \<in> fv (Fun g S)"
using assms
proof (induction T arbitrary: S)
case (Cons t T S')
then obtain s S where S': "S' = s#S" by (cases S') simp_all
thus ?case using Cons by fastforce
qed simp
lemma finite_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[simp]: "finite (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s x)" by auto
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_Nil[simp]: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s [] = {}" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_singleton[simp]: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s [(t,s)] = fv t \<union> fv s" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_Cons: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((s,t)#F) = fv s \<union> fv t \<union> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<union> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by simp
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono: "set M \<subseteq> set N \<Longrightarrow> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s M \<subseteq> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s N" by auto
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI[intro]:
"f \<in> set F \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r f \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"f \<in> set F \<Longrightarrow> x \<in> fv (fst f) \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"f \<in> set F \<Longrightarrow> x \<in> fv (snd f) \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"(t,s) \<in> set F \<Longrightarrow> x \<in> fv t \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
"(t,s) \<in> set F \<Longrightarrow> x \<in> fv s \<Longrightarrow> x \<in> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"
using UN_I by fastforce+
lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_cons_subset: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \<subseteq> fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F)"
by auto
subsection \<open>Other lemmata\<close>
lemma nonvar_term_has_composed_shallow_term:
fixes t::"('f,'v) term"
assumes "\<not>(\<exists>x. t = Var x)"
shows "\<exists>f T. Fun f T \<sqsubseteq> t \<and> (\<forall>s \<in> set T. (\<exists>c. s = Fun c []) \<or> (\<exists>x. s = Var x))"
proof -
let ?Q = "\<lambda>S. \<forall>s \<in> set S. (\<exists>c. s = Fun c []) \<or> (\<exists>x. s = Var x)"
let ?P = "\<lambda>t. \<exists>g S. Fun g S \<sqsubseteq> t \<and> ?Q S"
{ fix t::"('f,'v) term"
have "(\<exists>x. t = Var x) \<or> ?P t"
proof (induction t)
case (Fun h R) show ?case
proof (cases "R = [] \<or> (\<forall>r \<in> set R. \<exists>x. r = Var x)")
case False
then obtain r g S where "r \<in> set R" "?P r" "Fun g S \<sqsubseteq> r" "?Q S" using Fun.IH by fast
thus ?thesis by auto
qed force
qed simp
} thus ?thesis using assms by blast
qed
end

View File

@ -0,0 +1,492 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Miscellaneous.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>Miscellaneous Lemmata\<close>
theory Miscellaneous
imports Main "HOL-Library.Sublist" "HOL-Library.While_Combinator"
begin
subsection \<open>List: zip, filter, map\<close>
lemma zip_arg_subterm_split:
assumes "(x,y) \<in> set (zip xs ys)"
obtains xs' xs'' ys' ys'' where "xs = xs'@x#xs''" "ys = ys'@y#ys''" "length xs' = length ys'"
proof -
from assms have "\<exists>zs zs' vs vs'. xs = zs@x#zs' \<and> ys = vs@y#vs' \<and> length zs = length vs"
proof (induction ys arbitrary: xs)
case (Cons y' ys' xs)
then obtain x' xs' where x': "xs = x'#xs'"
by (metis empty_iff list.exhaust list.set(1) set_zip_leftD)
show ?case
by (cases "(x, y) \<in> set (zip xs' ys')",
metis \<open>xs = x'#xs'\<close> Cons.IH[of xs'] Cons_eq_appendI list.size(4),
use Cons.prems x' in fastforce)
qed simp
thus ?thesis using that by blast
qed
lemma zip_arg_index:
assumes "(x,y) \<in> set (zip xs ys)"
obtains i where "xs ! i = x" "ys ! i = y" "i < length xs" "i < length ys"
proof -
obtain xs1 xs2 ys1 ys2 where "xs = xs1@x#xs2" "ys = ys1@y#ys2" "length xs1 = length ys1"
using zip_arg_subterm_split[OF assms] by moura
thus ?thesis using nth_append_length[of xs1 x xs2] nth_append_length[of ys1 y ys2] that by simp
qed
lemma filter_nth: "i < length (filter P xs) \<Longrightarrow> P (filter P xs ! i)"
using nth_mem by force
lemma list_all_filter_eq: "list_all P xs \<Longrightarrow> filter P xs = xs"
by (metis list_all_iff filter_True)
lemma list_all_filter_nil:
assumes "list_all P xs"
and "\<And>x. P x \<Longrightarrow> \<not>Q x"
shows "filter Q xs = []"
using assms by (induct xs) simp_all
lemma list_all_concat: "list_all (list_all f) P \<longleftrightarrow> list_all f (concat P)"
by (induct P) auto
lemma map_upt_index_eq:
assumes "j < length xs"
shows "(map (\<lambda>i. xs ! is i) [0..<length xs]) ! j = xs ! is j"
using assms by (simp add: map_nth)
lemma map_snd_list_insert_distrib:
assumes "\<forall>(i,p) \<in> insert x (set xs). \<forall>(i',p') \<in> insert x (set xs). p = p' \<longrightarrow> i = i'"
shows "map snd (List.insert x xs) = List.insert (snd x) (map snd xs)"
using assms
proof (induction xs rule: List.rev_induct)
case (snoc y xs)
hence IH: "map snd (List.insert x xs) = List.insert (snd x) (map snd xs)" by fastforce
obtain iy py where y: "y = (iy,py)" by (metis surj_pair)
obtain ix px where x: "x = (ix,px)" by (metis surj_pair)
have "(ix,px) \<in> insert x (set (y#xs))" "(iy,py) \<in> insert x (set (y#xs))" using y x by auto
hence *: "iy = ix" when "py = px" using that snoc.prems by auto
show ?case
proof (cases "px = py")
case True
hence "y = x" using * y x by auto
thus ?thesis using IH by simp
next
case False
hence "y \<noteq> x" using y x by simp
have "List.insert x (xs@[y]) = (List.insert x xs)@[y]"
proof -
have 1: "insert y (set xs) = set (xs@[y])" by simp
have 2: "x \<notin> insert y (set xs) \<or> x \<in> set xs" using \<open>y \<noteq> x\<close> by blast
show ?thesis using 1 2 by (metis (no_types) List.insert_def append_Cons insertCI)
qed
thus ?thesis using IH y x False by (auto simp add: List.insert_def)
qed
qed simp
lemma map_append_inv: "map f xs = ys@zs \<Longrightarrow> \<exists>vs ws. xs = vs@ws \<and> map f vs = ys \<and> map f ws = zs"
proof (induction xs arbitrary: ys zs)
case (Cons x xs')
note prems = Cons.prems
note IH = Cons.IH
show ?case
proof (cases ys)
case (Cons y ys')
then obtain vs' ws where *: "xs' = vs'@ws" "map f vs' = ys'" "map f ws = zs"
using prems IH[of ys' zs] by auto
hence "x#xs' = (x#vs')@ws" "map f (x#vs') = y#ys'" using Cons prems by force+
thus ?thesis by (metis Cons *(3))
qed (use prems in simp)
qed simp
subsection \<open>List: subsequences\<close>
lemma subseqs_set_subset:
assumes "ys \<in> set (subseqs xs)"
shows "set ys \<subseteq> set xs"
using assms subseqs_powset[of xs] by auto
lemma subset_sublist_exists:
"ys \<subseteq> set xs \<Longrightarrow> \<exists>zs. set zs = ys \<and> zs \<in> set (subseqs xs)"
proof (induction xs arbitrary: ys)
case Cons thus ?case by (metis (no_types, lifting) Pow_iff imageE subseqs_powset)
qed simp
lemma map_subseqs: "map (map f) (subseqs xs) = subseqs (map f xs)"
proof (induct xs)
case (Cons x xs)
have "map (Cons (f x)) (map (map f) (subseqs xs)) = map (map f) (map (Cons x) (subseqs xs))"
by (induct "subseqs xs") auto
thus ?case by (simp add: Let_def Cons)
qed simp
lemma subseqs_Cons:
assumes "ys \<in> set (subseqs xs)"
shows "ys \<in> set (subseqs (x#xs))"
by (metis assms Un_iff set_append subseqs.simps(2))
lemma subseqs_subset:
assumes "ys \<in> set (subseqs xs)"
shows "set ys \<subseteq> set xs"
using assms by (metis Pow_iff image_eqI subseqs_powset)
subsection \<open>List: prefixes, suffixes\<close>
lemma suffix_Cons': "suffix [x] (y#ys) \<Longrightarrow> suffix [x] ys \<or> (y = x \<and> ys = [])"
using suffix_Cons[of "[x]"] by auto
lemma prefix_Cons': "prefix (x#xs) (x#ys) \<Longrightarrow> prefix xs ys"
by simp
lemma prefix_map: "prefix xs (map f ys) \<Longrightarrow> \<exists>zs. prefix zs ys \<and> map f zs = xs"
using map_append_inv unfolding prefix_def by fast
lemma length_prefix_ex:
assumes "n \<le> length xs"
shows "\<exists>ys zs. xs = ys@zs \<and> length ys = n"
using assms
proof (induction n)
case (Suc n)
then obtain ys zs where IH: "xs = ys@zs" "length ys = n" by moura
hence "length zs > 0" using Suc.prems(1) by auto
then obtain v vs where v: "zs = v#vs" by (metis Suc_length_conv gr0_conv_Suc)
hence "length (ys@[v]) = Suc n" using IH(2) by simp
thus ?case using IH(1) v by (metis append.assoc append_Cons append_Nil)
qed simp
lemma length_prefix_ex':
assumes "n < length xs"
shows "\<exists>ys zs. xs = ys@xs ! n#zs \<and> length ys = n"
proof -
obtain ys zs where xs: "xs = ys@zs" "length ys = n" using assms length_prefix_ex[of n xs] by moura
hence "length zs > 0" using assms by auto
then obtain v vs where v: "zs = v#vs" by (metis Suc_length_conv gr0_conv_Suc)
hence "(ys@zs) ! n = v" using xs by auto
thus ?thesis using v xs by auto
qed
lemma length_prefix_ex2:
assumes "i < length xs" "j < length xs" "i < j"
shows "\<exists>ys zs vs. xs = ys@xs ! i#zs@xs ! j#vs \<and> length ys = i \<and> length zs = j - i - 1"
by (smt assms length_prefix_ex' nth_append append.assoc append.simps(2) add_diff_cancel_left'
diff_Suc_1 length_Cons length_append)
subsection \<open>List: products\<close>
lemma product_lists_Cons:
"x#xs \<in> set (product_lists (y#ys)) \<longleftrightarrow> (xs \<in> set (product_lists ys) \<and> x \<in> set y)"
by auto
lemma product_lists_in_set_nth:
assumes "xs \<in> set (product_lists ys)"
shows "\<forall>i<length ys. xs ! i \<in> set (ys ! i)"
proof -
have 0: "length ys = length xs" using assms(1) by (simp add: in_set_product_lists_length)
thus ?thesis using assms
proof (induction ys arbitrary: xs)
case (Cons y ys)
obtain x xs' where xs: "xs = x#xs'" using Cons.prems(1) by (metis length_Suc_conv)
hence "xs' \<in> set (product_lists ys) \<Longrightarrow> \<forall>i<length ys. xs' ! i \<in> set (ys ! i)"
"length ys = length xs'" "x#xs' \<in> set (product_lists (y#ys))"
using Cons by simp_all
thus ?case using xs product_lists_Cons[of x xs' y ys] by (simp add: nth_Cons')
qed simp
qed
lemma product_lists_in_set_nth':
assumes "\<forall>i<length xs. ys ! i \<in> set (xs ! i)"
and "length xs = length ys"
shows "ys \<in> set (product_lists xs)"
using assms
proof (induction xs arbitrary: ys)
case (Cons x xs)
obtain y ys' where ys: "ys = y#ys'" using Cons.prems(2) by (metis length_Suc_conv)
hence "ys' \<in> set (product_lists xs)" "y \<in> set x" "length xs = length ys'"
using Cons by fastforce+
thus ?case using ys product_lists_Cons[of y ys' x xs] by (simp add: nth_Cons')
qed simp
subsection \<open>Other Lemmata\<close>
lemma inv_set_fset: "finite M \<Longrightarrow> set (inv set M) = M"
unfolding inv_def by (metis (mono_tags) finite_list someI_ex)
lemma lfp_eqI':
assumes "mono f"
and "f C = C"
and "\<forall>X \<in> Pow C. f X = X \<longrightarrow> X = C"
shows "lfp f = C"
by (metis PowI assms lfp_lowerbound lfp_unfold subset_refl)
lemma lfp_while':
fixes f::"'a set \<Rightarrow> 'a set" and M::"'a set"
defines "N \<equiv> while (\<lambda>A. f A \<noteq> A) f {}"
assumes f_mono: "mono f"
and N_finite: "finite N"
and N_supset: "f N \<subseteq> N"
shows "lfp f = N"
proof -
have *: "f X \<subseteq> N" when "X \<subseteq> N" for X using N_supset monoD[OF f_mono that] by blast
show ?thesis
using lfp_while[OF f_mono * N_finite]
by (simp add: N_def)
qed
lemma lfp_while'':
fixes f::"'a set \<Rightarrow> 'a set" and M::"'a set"
defines "N \<equiv> while (\<lambda>A. f A \<noteq> A) f {}"
assumes f_mono: "mono f"
and lfp_finite: "finite (lfp f)"
shows "lfp f = N"
proof -
have *: "f X \<subseteq> lfp f" when "X \<subseteq> lfp f" for X
using lfp_fixpoint[OF f_mono] monoD[OF f_mono that]
by blast
show ?thesis
using lfp_while[OF f_mono * lfp_finite]
by (simp add: N_def)
qed
lemma preordered_finite_set_has_maxima:
assumes "finite A" "A \<noteq> {}"
shows "\<exists>a::'a::{preorder} \<in> A. \<forall>b \<in> A. \<not>(a < b)"
using assms
proof (induction A rule: finite_induct)
case (insert a A) thus ?case
by (cases "A = {}", simp, metis insert_iff order_trans less_le_not_le)
qed simp
lemma partition_index_bij:
fixes n::nat
obtains I k where
"bij_betw I {0..<n} {0..<n}" "k \<le> n"
"\<forall>i. i < k \<longrightarrow> P (I i)"
"\<forall>i. k \<le> i \<and> i < n \<longrightarrow> \<not>(P (I i))"
proof -
define A where "A = filter P [0..<n]"
define B where "B = filter (\<lambda>i. \<not>P i) [0..<n]"
define k where "k = length A"
define I where "I = (\<lambda>n. (A@B) ! n)"
note defs = A_def B_def k_def I_def
have k1: "k \<le> n" by (metis defs(1,3) diff_le_self dual_order.trans length_filter_le length_upt)
have "i < k \<Longrightarrow> P (A ! i)" for i by (metis defs(1,3) filter_nth)
hence k2: "i < k \<Longrightarrow> P ((A@B) ! i)" for i by (simp add: defs nth_append)
have "i < length B \<Longrightarrow> \<not>(P (B ! i))" for i by (metis defs(2) filter_nth)
hence "i < length B \<Longrightarrow> \<not>(P ((A@B) ! (k + i)))" for i using k_def by simp
hence "k \<le> i \<and> i < k + length B \<Longrightarrow> \<not>(P ((A@B) ! i))" for i
by (metis add.commute add_less_imp_less_right le_add_diff_inverse2)
hence k3: "k \<le> i \<and> i < n \<Longrightarrow> \<not>(P ((A@B) ! i))" for i by (simp add: defs sum_length_filter_compl)
have *: "length (A@B) = n" "set (A@B) = {0..<n}" "distinct (A@B)"
by (metis defs(1,2) diff_zero length_append length_upt sum_length_filter_compl)
(auto simp add: defs)
have I: "bij_betw I {0..<n} {0..<n}"
proof (intro bij_betwI')
fix x y show "x \<in> {0..<n} \<Longrightarrow> y \<in> {0..<n} \<Longrightarrow> (I x = I y) = (x = y)"
by (metis *(1,3) defs(4) nth_eq_iff_index_eq atLeastLessThan_iff)
next
fix x show "x \<in> {0..<n} \<Longrightarrow> I x \<in> {0..<n}"
by (metis *(1,2) defs(4) atLeastLessThan_iff nth_mem)
next
fix y show "y \<in> {0..<n} \<Longrightarrow> \<exists>x \<in> {0..<n}. y = I x"
by (metis * defs(4) atLeast0LessThan distinct_Ex1 lessThan_iff)
qed
show ?thesis using k1 k2 k3 I that by (simp add: defs)
qed
lemma finite_lists_length_eq':
assumes "\<And>x. x \<in> set xs \<Longrightarrow> finite {y. P x y}"
shows "finite {ys. length xs = length ys \<and> (\<forall>y \<in> set ys. \<exists>x \<in> set xs. P x y)}"
proof -
define Q where "Q \<equiv> \<lambda>ys. \<forall>y \<in> set ys. \<exists>x \<in> set xs. P x y"
define M where "M \<equiv> {y. \<exists>x \<in> set xs. P x y}"
have 0: "finite M" using assms unfolding M_def by fastforce
have "Q ys \<longleftrightarrow> set ys \<subseteq> M"
"(Q ys \<and> length ys = length xs) \<longleftrightarrow> (length xs = length ys \<and> Q ys)"
for ys
unfolding Q_def M_def by auto
thus ?thesis
using finite_lists_length_eq[OF 0, of "length xs"]
unfolding Q_def by presburger
qed
lemma trancl_eqI:
assumes "\<forall>(a,b) \<in> A. \<forall>(c,d) \<in> A. b = c \<longrightarrow> (a,d) \<in> A"
shows "A = A\<^sup>+"
proof
show "A\<^sup>+ \<subseteq> A"
proof
fix x assume x: "x \<in> A\<^sup>+"
then obtain a b where ab: "x = (a,b)" by (metis surj_pair)
hence "(a,b) \<in> A\<^sup>+" using x by metis
hence "(a,b) \<in> A" using assms by (induct rule: trancl_induct) auto
thus "x \<in> A" using ab by metis
qed
qed auto
lemma trancl_eqI':
assumes "\<forall>(a,b) \<in> A. \<forall>(c,d) \<in> A. b = c \<and> a \<noteq> d \<longrightarrow> (a,d) \<in> A"
and "\<forall>(a,b) \<in> A. a \<noteq> b"
shows "A = {(a,b) \<in> A\<^sup>+. a \<noteq> b}"
proof
show "{(a,b) \<in> A\<^sup>+. a \<noteq> b} \<subseteq> A"
proof
fix x assume x: "x \<in> {(a,b) \<in> A\<^sup>+. a \<noteq> b}"
then obtain a b where ab: "x = (a,b)" by (metis surj_pair)
hence "(a,b) \<in> A\<^sup>+" "a \<noteq> b" using x by blast+
hence "(a,b) \<in> A"
proof (induction rule: trancl_induct)
case base thus ?case by blast
next
case step thus ?case using assms(1) by force
qed
thus "x \<in> A" using ab by metis
qed
qed (use assms(2) in auto)
lemma distinct_concat_idx_disjoint:
assumes xs: "distinct (concat xs)"
and ij: "i < length xs" "j < length xs" "i < j"
shows "set (xs ! i) \<inter> set (xs ! j) = {}"
proof -
obtain ys zs vs where ys: "xs = ys@xs ! i#zs@xs ! j#vs" "length ys = i" "length zs = j - i - 1"
using length_prefix_ex2[OF ij] by moura
thus ?thesis
using xs concat_append[of "ys@xs ! i#zs" "xs ! j#vs"]
distinct_append[of "concat (ys@xs ! i#zs)" "concat (xs ! j#vs)"]
by auto
qed
lemma remdups_ex2:
"length (remdups xs) > 1 \<Longrightarrow> \<exists>a \<in> set xs. \<exists>b \<in> set xs. a \<noteq> b"
by (metis distinct_Ex1 distinct_remdups less_trans nth_mem set_remdups zero_less_one zero_neq_one)
lemma trancl_minus_refl_idem:
defines "cl \<equiv> \<lambda>ts. {(a,b) \<in> ts\<^sup>+. a \<noteq> b}"
shows "cl (cl ts) = cl ts"
proof -
have 0: "(ts\<^sup>+)\<^sup>+ = ts\<^sup>+" "cl ts \<subseteq> ts\<^sup>+" "(cl ts)\<^sup>+ \<subseteq> (ts\<^sup>+)\<^sup>+"
proof -
show "(ts\<^sup>+)\<^sup>+ = ts\<^sup>+" "cl ts \<subseteq> ts\<^sup>+" unfolding cl_def by auto
thus "(cl ts)\<^sup>+ \<subseteq> (ts\<^sup>+)\<^sup>+" using trancl_mono[of _ "cl ts" "ts\<^sup>+"] by blast
qed
have 1: "t \<in> cl (cl ts)" when t: "t \<in> cl ts" for t
using t 0 unfolding cl_def by fast
have 2: "t \<in> cl ts" when t: "t \<in> cl (cl ts)" for t
proof -
obtain a b where ab: "t = (a,b)" by (metis surj_pair)
have "t \<in> (cl ts)\<^sup>+" and a_neq_b: "a \<noteq> b" using t unfolding cl_def ab by force+
hence "t \<in> ts\<^sup>+" using 0 by blast
thus ?thesis using a_neq_b unfolding cl_def ab by blast
qed
show ?thesis using 1 2 by blast
qed
subsection \<open>Infinite Paths in Relations as Mappings from Naturals to States\<close>
context
begin
private fun rel_chain_fun::"nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a" where
"rel_chain_fun 0 x _ _ = x"
| "rel_chain_fun (Suc i) x y r = (if i = 0 then y else SOME z. (rel_chain_fun i x y r, z) \<in> r)"
lemma infinite_chain_intro:
fixes r::"('a \<times> 'a) set"
assumes "\<forall>(a,b) \<in> r. \<exists>c. (b,c) \<in> r" "r \<noteq> {}"
shows "\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> r"
proof -
from assms(2) obtain a b where "(a,b) \<in> r" by auto
let ?P = "\<lambda>i. (rel_chain_fun i a b r, rel_chain_fun (Suc i) a b r) \<in> r"
let ?Q = "\<lambda>i. \<exists>z. (rel_chain_fun i a b r, z) \<in> r"
have base: "?P 0" using \<open>(a,b) \<in> r\<close> by auto
have step: "?P (Suc i)" when i: "?P i" for i
proof -
have "?Q (Suc i)" using assms(1) i by auto
thus ?thesis using someI_ex[OF \<open>?Q (Suc i)\<close>] by auto
qed
have "\<forall>i::nat. (rel_chain_fun i a b r, rel_chain_fun (Suc i) a b r) \<in> r"
using base step nat_induct[of ?P] by simp
thus ?thesis by fastforce
qed
end
lemma infinite_chain_intro':
fixes r::"('a \<times> 'a) set"
assumes base: "\<exists>b. (x,b) \<in> r" and step: "\<forall>b. (x,b) \<in> r\<^sup>+ \<longrightarrow> (\<exists>c. (b,c) \<in> r)"
shows "\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> r"
proof -
let ?s = "{(a,b) \<in> r. a = x \<or> (x,a) \<in> r\<^sup>+}"
have "?s \<noteq> {}" using base by auto
have "\<exists>c. (b,c) \<in> ?s" when ab: "(a,b) \<in> ?s" for a b
proof (cases "a = x")
case False
hence "(x,a) \<in> r\<^sup>+" using ab by auto
hence "(x,b) \<in> r\<^sup>+" using \<open>(a,b) \<in> ?s\<close> by auto
thus ?thesis using step by auto
qed (use ab step in auto)
hence "\<exists>f. \<forall>i. (f i, f (Suc i)) \<in> ?s" using infinite_chain_intro[of ?s] \<open>?s \<noteq> {}\<close> by blast
thus ?thesis by auto
qed
lemma infinite_chain_mono:
assumes "S \<subseteq> T" "\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> S"
shows "\<exists>f. \<forall>i::nat. (f i, f (Suc i)) \<in> T"
using assms by auto
end

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,13 @@
chapter AFP
session "Stateful_Protocol_Composition_and_Typing-devel" (AFP) = "First_Order_Terms" +
options [timeout = 2400]
directories
"examples"
theories
"Stateful_Compositionality"
"Examples"
document_files
"root.tex"
"root.bib"

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,47 @@
@InProceedings{ hess.ea:formalizing:2017,
author = {Andreas V. Hess and Sebastian M{\"{o}}dersheim},
title = {{Formalizing and Proving a Typing Result for Security Protocols in Isabelle/HOL}},
booktitle = {30th {IEEE} Computer Security Foundations Symposium, {CSF} 2017, Santa Barbara, CA, USA, August
21-25, 2017},
pages = {451--463},
publisher = {{IEEE} Computer Society},
year = 2017,
doi = {10.1109/CSF.2017.27}
}
@InProceedings{ hess.ea:typing:2018,
author = {Andreas V. Hess and Sebastian M{\"{o}}dersheim},
title = {{A Typing Result for Stateful Protocols}},
booktitle = {31st {IEEE} Computer Security Foundations Symposium, {CSF} 2018, Oxford, United Kingdom, July 9-12,
2018},
pages = {374--388},
publisher = {{IEEE} Computer Society},
year = 2018,
doi = {10.1109/CSF.2018.00034}
}
@InProceedings{ hess.ea:stateful:2018,
author = {Andreas V. Hess and Sebastian M{\"{o}}dersheim and Achim D. Brucker},
editor = {Javier L{\'{o}}pez and Jianying Zhou and Miguel Soriano},
title = {{Stateful Protocol Composition}},
booktitle = {Computer Security - 23rd European Symposium on Research in Computer Security, {ESORICS} 2018,
Barcelona, Spain, September 3-7, 2018, Proceedings, Part {I}},
series = {Lecture Notes in Computer Science},
volume = 11098,
pages = {427--446},
publisher = {Springer},
year = 2018,
doi = {10.1007/978-3-319-99073-6_21}
}
@PhDThesis{ hess:typing:2018,
author = {Andreas Viktor Hess},
title = {Typing and Compositionality for Stateful Security Protocols},
year = {2019},
url = {https://orbit.dtu.dk/en/publications/typing-and-compositionality-for-stateful-security-protocols},
language = {English},
series = {TU Compute PHD-2018},
publisher = {DTU Compute}
}

View File

@ -0,0 +1,151 @@
\documentclass[10pt,DIV16,a4paper,abstract=true,twoside=semi,openright]
{scrreprt}
\usepackage[USenglish]{babel}
\usepackage[numbers, sort&compress]{natbib}
\usepackage{isabelle,isabellesym}
\usepackage{booktabs}
\usepackage{paralist}
\usepackage{graphicx}
\usepackage{amssymb}
\usepackage{xspace}
\usepackage{xcolor}
\usepackage{hyperref}
\sloppy
\pagestyle{headings}
\isabellestyle{default}
\setcounter{tocdepth}{1}
\newcommand{\ie}{i.\,e.\xspace}
\newcommand{\eg}{e.\,g.\xspace}
\newcommand{\thy}{\isabellecontext}
\renewcommand{\isamarkupsection}[1]{%
\begingroup%
\def\isacharunderscore{\textunderscore}%
\section{#1 (\thy)}%
\def\isacharunderscore{-}%
\expandafter\label{sec:\isabellecontext}%
\endgroup%
}
\title{Stateful Protocol Composition and Typing}
\author{%
\href{https://www.dtu.dk/english/service/phonebook/person?id=64207}{Andreas~V.~Hess}\footnotemark[1]
\and
\href{https://people.compute.dtu.dk/samo/}{Sebastian~M{\"o}dersheim}\footnotemark[1]
\and
\href{http://www.brucker.ch/}{Achim~D.~Brucker}\footnotemark[2]
}
\publishers{%
\footnotemark[1]~DTU Compute, Technical University of Denmark, Lyngby, Denmark\texorpdfstring{\\}{, }
\texttt{\{avhe, samo\}@dtu.dk}\\[2em]
%
\footnotemark[2]~
Department of Computer Science, University of Exeter, Exeter, UK\texorpdfstring{\\}{, }
\texttt{a.brucker@exeter.ac.uk}
%
}
\begin{document}
\maketitle
\begin{abstract}
\begin{quote}
We provide in this AFP entry several relative soundness results for security protocols.
In particular, we prove typing and compositionality results for stateful protocols (i.e., protocols with mutable state that may span several sessions), and that focuses on reachability properties.
Such results are useful to simplify protocol verification by reducing it to a simpler problem: Typing results give conditions under which it is safe to verify a protocol in a typed model where only ``well-typed'' attacks can occur whereas compositionality results allow us to verify a composed protocol by only verifying the component protocols in isolation.
The conditions on the protocols under which the results hold are furthermore syntactic in nature allowing for full automation.
The foundation presented here is used in another entry to provide fully automated and formalized security proofs of stateful protocols.
\bigskip
\noindent{\textbf{Keywords:}}
Security protocols, stateful protocols, relative soundness results, proof assistants, Isabelle/HOL, compositionality
\end{quote}
\end{abstract}
\tableofcontents
\cleardoublepage
\chapter{Introduction}
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}.
\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}.
We introduce our compositionality results in \autoref{cha:composition} and \autoref{cha:stateful-composition}.
Finally, we present two example protocols \autoref{cha:examples}.
\paragraph{Acknowledgments}
This work was supported by the Sapere-Aude project ``Composec: Secure Composition of Distributed Systems'', grant 4184-00334B of the Danish Council for Independent Research.
\clearpage
\begin{figure}
\centering
\includegraphics[height=\textheight]{session_graph}
\caption{The Dependency Graph of the Isabelle Theories.\label{fig:session-graph}}
\end{figure}
\clearpage
% \input{session}
\chapter{Preliminaries and Intruder Model}
\label{cha:preliminaries}
In this chapter, we introduce the formal preliminaries, including the intruder model and related lemmata.
\input{Miscellaneous.tex}
\input{Messages.tex}
\input{More_Unification.tex}
\input{Intruder_Deduction.tex}
\chapter{The Typing Result for Non-Stateful Protocols}
\label{cha:typing}
In this chapter, we formalize and prove a typing result for ``stateless'' security protocols.
This work is described in more detail in~\cite{hess.ea:formalizing:2017} and~\cite[chapter 3]{hess:typing:2018}.
\input{Strands_and_Constraints.tex}
\input{Lazy_Intruder.tex}
\input{Typed_Model.tex}
\input{Typing_Result.tex}
\chapter{The Typing Result for Stateful Protocols}
\label{cha:stateful-typing}
In this chapter, we lift the typing result to stateful protocols.
For more details, we refer the reader to~\cite{hess.ea:typing:2018} and~\cite[chapter 4]{hess:typing:2018}.
\input{Stateful_Strands.tex}
\input{Stateful_Typing.tex}
\chapter{The Parallel Composition Result for Non-Stateful Protocols}
\label{cha:composition}
In this chapter, we formalize and prove a compositionality result for security protocols.
This work is an extension of the work described in~\cite{hess.ea:stateful:2018} and~\cite[chapter 5]{hess:typing:2018}.
\input{Labeled_Strands.tex}
\input{Parallel_Compositionality.tex}
\chapter{The Stateful Protocol Composition Result}
\label{cha:stateful-composition}
In this chapter, we extend the compositionality result to stateful security protocols.
This work is an extension of the work described in~\cite{hess.ea:stateful:2018} and~\cite[chapter 5]{hess:typing:2018}.
\input{Labeled_Stateful_Strands.tex}
\input{Stateful_Compositionality.tex}
\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.
\input{Example_TLS.tex}
\input{Example_Keyserver.tex}
{\small
\bibliographystyle{abbrvnat}
\bibliography{root}
}
\end{document}
%%% Local Variables:
%%% mode: latex
%%% TeX-master: t
%%% End:

View File

@ -0,0 +1,404 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Example_Keyserver.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>The Keyserver Example\<close>
theory Example_Keyserver
imports "../Stateful_Compositionality"
begin
declare [[code_timing]]
subsection \<open>Setup\<close>
subsubsection \<open>Datatypes and functions setup\<close>
datatype ex_lbl = Label1 ("\<one>") | Label2 ("\<two>")
datatype ex_atom =
Agent | Value | Attack | PrivFunSec
| Bot
datatype ex_fun =
ring | valid | revoked | events | beginauth nat | endauth nat | pubkeys | seen
| invkey | tuple | tuple' | attack nat
| sign | crypt | update | pw
| encodingsecret | pubkey nat
| pubconst ex_atom nat
type_synonym ex_type = "(ex_fun, ex_atom) term_type"
type_synonym ex_var = "ex_type \<times> nat"
lemma ex_atom_UNIV:
"(UNIV::ex_atom set) = {Agent, Value, Attack, PrivFunSec, Bot}"
by (auto intro: ex_atom.exhaust)
instance ex_atom::finite
by intro_classes (metis ex_atom_UNIV finite.emptyI finite.insertI)
lemma ex_lbl_UNIV:
"(UNIV::ex_lbl set) = {Label1, Label2}"
by (auto intro: ex_lbl.exhaust)
type_synonym ex_term = "(ex_fun, ex_var) term"
type_synonym ex_terms = "(ex_fun, ex_var) terms"
primrec arity::"ex_fun \<Rightarrow> nat" where
"arity ring = 2"
| "arity valid = 3"
| "arity revoked = 3"
| "arity events = 1"
| "arity (beginauth _) = 3"
| "arity (endauth _) = 3"
| "arity pubkeys = 2"
| "arity seen = 2"
| "arity invkey = 2"
| "arity tuple = 2"
| "arity tuple' = 2"
| "arity (attack _) = 0"
| "arity sign = 2"
| "arity crypt = 2"
| "arity update = 4"
| "arity pw = 2"
| "arity (pubkey _) = 0"
| "arity encodingsecret = 0"
| "arity (pubconst _ _) = 0"
fun public::"ex_fun \<Rightarrow> bool" where
"public (pubkey _) = False"
| "public encodingsecret = False"
| "public _ = True"
fun Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t::"ex_term list \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t [k,m] = ([Fun invkey [Fun encodingsecret [], k]], [m])"
| "Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t _ = ([], [])"
fun Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n::"ex_term list \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n [k,m] = ([], [m])"
| "Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n _ = ([], [])"
fun Ana::"ex_term \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana (Fun tuple T) = ([], T)"
| "Ana (Fun tuple' T) = ([], T)"
| "Ana (Fun sign T) = Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n T"
| "Ana (Fun crypt T) = Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t T"
| "Ana _ = ([], [])"
subsubsection \<open>Keyserver example: Locale interpretation\<close>
lemma assm1:
"Ana t = (K,M) \<Longrightarrow> fv\<^sub>s\<^sub>e\<^sub>t (set K) \<subseteq> fv t"
"Ana t = (K,M) \<Longrightarrow> (\<And>g S'. Fun g S' \<sqsubseteq> t \<Longrightarrow> length S' = arity g)
\<Longrightarrow> k \<in> set K \<Longrightarrow> Fun f T' \<sqsubseteq> k \<Longrightarrow> length T' = arity f"
"Ana t = (K,M) \<Longrightarrow> K \<noteq> [] \<or> M \<noteq> [] \<Longrightarrow> Ana (t \<cdot> \<delta>) = (K \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>, M \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>)"
by (rule Ana.cases[of "t"], auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)+
lemma assm2: "Ana (Fun f T) = (K, M) \<Longrightarrow> set M \<subseteq> set T"
by (rule Ana.cases[of "Fun f T"]) (auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)
lemma assm6: "0 < arity f \<Longrightarrow> public f" by (cases f) simp_all
global_interpretation im: intruder_model arity public Ana
defines wf\<^sub>t\<^sub>r\<^sub>m = "im.wf\<^sub>t\<^sub>r\<^sub>m"
by unfold_locales (metis assm1(1), metis assm1(2),rule Ana.simps, metis assm2, metis assm1(3))
type_synonym ex_strand_step = "(ex_fun,ex_var) strand_step"
type_synonym ex_strand = "(ex_fun,ex_var) strand"
subsubsection \<open>Typing function\<close>
definition \<Gamma>\<^sub>v::"ex_var \<Rightarrow> ex_type" where
"\<Gamma>\<^sub>v v = (if (\<forall>t \<in> subterms (fst v). case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True)
then fst v else TAtom Bot)"
fun \<Gamma>::"ex_term \<Rightarrow> ex_type" where
"\<Gamma> (Var v) = \<Gamma>\<^sub>v v"
| "\<Gamma> (Fun (attack _) _) = TAtom Attack"
| "\<Gamma> (Fun (pubkey _) _) = TAtom Value"
| "\<Gamma> (Fun encodingsecret _) = TAtom PrivFunSec"
| "\<Gamma> (Fun (pubconst \<tau> _) _) = TAtom \<tau>"
| "\<Gamma> (Fun f T) = TComp f (map \<Gamma> T)"
subsubsection \<open>Locale interpretation: typed model\<close>
lemma assm7: "arity c = 0 \<Longrightarrow> \<exists>a. \<forall>X. \<Gamma> (Fun c X) = TAtom a" by (cases c) simp_all
lemma assm8: "0 < arity f \<Longrightarrow> \<Gamma> (Fun f X) = TComp f (map \<Gamma> X)" by (cases f) simp_all
lemma assm9: "infinite {c. \<Gamma> (Fun c []) = TAtom a \<and> public c}"
proof -
let ?T = "(range (pubconst a))::ex_fun set"
have *:
"\<And>x y::nat. x \<in> UNIV \<Longrightarrow> y \<in> UNIV \<Longrightarrow> (pubconst a x = pubconst a y) = (x = y)"
"\<And>x::nat. x \<in> UNIV \<Longrightarrow> pubconst a x \<in> ?T"
"\<And>y::ex_fun. y \<in> ?T \<Longrightarrow> \<exists>x \<in> UNIV. y = pubconst a x"
by auto
have "?T \<subseteq> {c. \<Gamma> (Fun c []) = TAtom a \<and> public c}" by auto
moreover have "\<exists>f::nat \<Rightarrow> ex_fun. bij_betw f UNIV ?T"
using bij_betwI'[OF *] by blast
hence "infinite ?T" by (metis nat_not_finite bij_betw_finite)
ultimately show ?thesis using infinite_super by blast
qed
lemma assm10: "TComp f T \<sqsubseteq> \<Gamma> t \<Longrightarrow> arity f > 0"
proof (induction rule: \<Gamma>.induct)
case (1 x)
hence *: "TComp f T \<sqsubseteq> \<Gamma>\<^sub>v x" by simp
hence "\<Gamma>\<^sub>v x \<noteq> TAtom Bot" unfolding \<Gamma>\<^sub>v_def by force
hence "\<forall>t \<in> subterms (fst x). case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True"
unfolding \<Gamma>\<^sub>v_def by argo
thus ?case using * unfolding \<Gamma>\<^sub>v_def by fastforce
qed auto
lemma assm11: "im.wf\<^sub>t\<^sub>r\<^sub>m (\<Gamma> (Var x))"
proof -
have "im.wf\<^sub>t\<^sub>r\<^sub>m (\<Gamma>\<^sub>v x)" unfolding \<Gamma>\<^sub>v_def im.wf\<^sub>t\<^sub>r\<^sub>m_def by auto
thus ?thesis by simp
qed
lemma assm12: "\<Gamma> (Var (\<tau>, n)) = \<Gamma> (Var (\<tau>, m))"
apply (cases "\<forall>t \<in> subterms \<tau>. case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True")
by (auto simp add: \<Gamma>\<^sub>v_def)
lemma Ana_const: "arity c = 0 \<Longrightarrow> Ana (Fun c T) = ([], [])"
by (cases c) simp_all
lemma Ana_subst': "Ana (Fun f T) = (K,M) \<Longrightarrow> Ana (Fun f T \<cdot> \<delta>) = (K \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>,M \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>)"
by (cases f) (auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)
global_interpretation tm: typed_model' arity public Ana \<Gamma>
by (unfold_locales, unfold wf\<^sub>t\<^sub>r\<^sub>m_def[symmetric])
(metis assm7, metis assm8, metis assm9, metis assm10, metis assm11, metis assm6,
metis assm12, metis Ana_const, metis Ana_subst')
subsubsection \<open>Locale interpretation: labeled stateful typed model\<close>
global_interpretation stm: labeled_stateful_typed_model' arity public Ana \<Gamma> tuple \<one> \<two>
by standard (rule arity.simps, metis Ana_subst', metis assm12, metis Ana_const, simp)
type_synonym ex_stateful_strand_step = "(ex_fun,ex_var) stateful_strand_step"
type_synonym ex_stateful_strand = "(ex_fun,ex_var) stateful_strand"
type_synonym ex_labeled_stateful_strand_step =
"(ex_fun,ex_var,ex_lbl) labeled_stateful_strand_step"
type_synonym ex_labeled_stateful_strand =
"(ex_fun,ex_var,ex_lbl) labeled_stateful_strand"
subsection \<open>Theorem: Type-flaw resistance of the keyserver example from the CSF18 paper\<close>
abbreviation "PK n \<equiv> Var (TAtom Value,n)"
abbreviation "A n \<equiv> Var (TAtom Agent,n)"
abbreviation "X n \<equiv> (TAtom Agent,n)"
abbreviation "ringset t \<equiv> Fun ring [Fun encodingsecret [], t]"
abbreviation "validset t t' \<equiv> Fun valid [Fun encodingsecret [], t, t']"
abbreviation "revokedset t t' \<equiv> Fun revoked [Fun encodingsecret [], t, t']"
abbreviation "eventsset \<equiv> Fun events [Fun encodingsecret []]"
(* Note: We will use S\<^sub>k\<^sub>s as a constraint, but it actually represents all steps that might occur
in the protocol *)
abbreviation S\<^sub>k\<^sub>s::"(ex_fun,ex_var) stateful_strand_step list" where
"S\<^sub>k\<^sub>s \<equiv> [
insert\<langle>Fun (attack 0) [], eventsset\<rangle>,
delete\<langle>PK 0, validset (A 0) (A 0)\<rangle>,
\<forall>(TAtom Agent,0)\<langle>PK 0 not in revokedset (A 0) (A 0)\<rangle>,
\<forall>(TAtom Agent,0)\<langle>PK 0 not in validset (A 0) (A 0)\<rangle>,
insert\<langle>PK 0, validset (A 0) (A 0)\<rangle>,
insert\<langle>PK 0, ringset (A 0)\<rangle>,
insert\<langle>PK 0, revokedset (A 0) (A 0)\<rangle>,
select\<langle>PK 0, validset (A 0) (A 0)\<rangle>,
select\<langle>PK 0, ringset (A 0)\<rangle>,
receive\<langle>Fun invkey [Fun encodingsecret [], PK 0]\<rangle>,
receive\<langle>Fun sign [Fun invkey [Fun encodingsecret [], PK 0], Fun tuple' [A 0, PK 0]]\<rangle>,
send\<langle>Fun invkey [Fun encodingsecret [], PK 0]\<rangle>,
send\<langle>Fun sign [Fun invkey [Fun encodingsecret [], PK 0], Fun tuple' [A 0, PK 0]]\<rangle>
]"
theorem "stm.tfr\<^sub>s\<^sub>s\<^sub>t S\<^sub>k\<^sub>s"
proof -
let ?M = "concat (map subterms_list (trms_list\<^sub>s\<^sub>s\<^sub>t S\<^sub>k\<^sub>s@map (pair' tuple) (setops_list\<^sub>s\<^sub>s\<^sub>t S\<^sub>k\<^sub>s)))"
have "comp_tfr\<^sub>s\<^sub>s\<^sub>t arity Ana \<Gamma> tuple ?M S\<^sub>k\<^sub>s" by eval
thus ?thesis by (rule stm.tfr\<^sub>s\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>s\<^sub>t)
qed
subsection \<open>Theorem: Type-flaw resistance of the keyserver examples from the ESORICS18 paper\<close>
abbreviation "signmsg t t' \<equiv> Fun sign [t, t']"
abbreviation "cryptmsg t t' \<equiv> Fun crypt [t, t']"
abbreviation "invkeymsg t \<equiv> Fun invkey [Fun encodingsecret [], t]"
abbreviation "updatemsg a b c d \<equiv> Fun update [a,b,c,d]"
abbreviation "pwmsg t t' \<equiv> Fun pw [t, t']"
abbreviation "beginauthset n t t' \<equiv> Fun (beginauth n) [Fun encodingsecret [], t, t']"
abbreviation "endauthset n t t' \<equiv> Fun (endauth n) [Fun encodingsecret [], t, t']"
abbreviation "pubkeysset t \<equiv> Fun pubkeys [Fun encodingsecret [], t]"
abbreviation "seenset t \<equiv> Fun seen [Fun encodingsecret [], t]"
declare [[coercion "Var::ex_var \<Rightarrow> ex_term"]]
declare [[coercion_enabled]]
(* Note: S'\<^sub>k\<^sub>s contains the (slightly over-approximated) steps that can occur in the
reachable constraints of \<P>\<^sub>k\<^sub>s,\<^sub>1 and \<P>\<^sub>k\<^sub>s,\<^sub>2 modulo variable renaming *)
definition S'\<^sub>k\<^sub>s::"ex_labeled_stateful_strand_step list" where
"S'\<^sub>k\<^sub>s \<equiv> [
\<^cancel>\<open>constraint steps from the first protocol (duplicate steps are ignored)\<close>
\<^cancel>\<open>rule R^1_1\<close>
\<langle>\<one>, send\<langle>invkeymsg (PK 0)\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 in validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<one>, receive\<langle>Fun (attack 0) []\<rangle>\<rangle>,
\<^cancel>\<open>rule R^2_1\<close>
\<langle>\<one>, send\<langle>signmsg (invkeymsg (PK 0)) (Fun tuple' [A 0, PK 0])\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 in validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, \<forall>X 0, X 1\<langle>PK 0 not in validset (Var (X 0)) (Var (X 1))\<rangle>\<rangle>,
\<langle>\<one>, \<forall>X 0, X 1\<langle>PK 0 not in revokedset (Var (X 0)) (Var (X 1))\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 not in beginauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^3_1\<close>
\<langle>\<star>, \<langle>PK 0 in beginauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 in endauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^4_1\<close>
\<langle>\<star>, receive\<langle>PK 0\<rangle>\<rangle>,
\<langle>\<star>, receive\<langle>invkeymsg (PK 0)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^5_1\<close>
\<langle>\<one>, insert\<langle>PK 0, ringset (A 0)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, beginauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, endauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^6_1\<close>
\<langle>\<one>, select\<langle>PK 0, ringset (A 0)\<rangle>\<rangle>,
\<langle>\<one>, delete\<langle>PK 0, ringset (A 0)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^7_1\<close>
\<langle>\<star>, \<langle>PK 0 not in endauthset 0 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, delete\<langle>PK 0, validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<one>, insert\<langle>PK 0, revokedset (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^8_1\<close>
\<^cancel>\<open>nothing new\<close>
\<^cancel>\<open>rule R^9_1\<close>
\<langle>\<one>, send\<langle>PK 0\<rangle>\<rangle>,
\<^cancel>\<open>rule R^10_1\<close>
\<langle>\<one>, send\<langle>Fun (attack 0) []\<rangle>\<rangle>,
\<^cancel>\<open>constraint steps from the second protocol (duplicate steps are ignored)\<close>
\<^cancel>\<open>rule R^2_1\<close>
\<langle>\<two>, send\<langle>invkeymsg (PK 0)\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 in validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<two>, receive\<langle>Fun (attack 1) []\<rangle>\<rangle>,
\<^cancel>\<open>rule R^2_2\<close>
\<langle>\<two>, send\<langle>cryptmsg (PK 0) (updatemsg (A 0) (A 1) (PK 1) (pwmsg (A 0) (A 1)))\<rangle>\<rangle>,
\<langle>\<two>, select\<langle>PK 0, pubkeysset (A 0)\<rangle>\<rangle>,
\<langle>\<two>, \<forall>X 0\<langle>PK 0 not in pubkeysset (Var (X 0))\<rangle>\<rangle>,
\<langle>\<two>, \<forall>X 0\<langle>PK 0 not in seenset (Var (X 0))\<rangle>\<rangle>,
\<^cancel>\<open>rule R^3_2\<close>
\<langle>\<star>, \<langle>PK 0 in beginauthset 1 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, \<langle>PK 0 in endauthset 1 (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^4_2\<close>
\<langle>\<star>, receive\<langle>PK 0\<rangle>\<rangle>,
\<langle>\<star>, receive\<langle>invkeymsg (PK 0)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^5_2\<close>
\<langle>\<two>, select\<langle>PK 0, pubkeysset (A 0)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, beginauthset 1 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<two>, receive\<langle>cryptmsg (PK 0) (updatemsg (A 0) (A 1) (PK 1) (pwmsg (A 0) (A 1)))\<rangle>\<rangle>,
\<^cancel>\<open>rule R^6_2\<close>
\<langle>\<star>, \<langle>PK 0 not in endauthset 1 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, validset (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<star>, insert\<langle>PK 0, endauthset 1 (A 0) (A 1)\<rangle>\<rangle>,
\<langle>\<two>, insert\<langle>PK 0, seenset (A 0)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^7_2\<close>
\<langle>\<two>, receive\<langle>pwmsg (A 0) (A 1)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^8_2\<close>
\<^cancel>\<open>nothing new\<close>
\<^cancel>\<open>rule R^9_2\<close>
\<langle>\<two>, insert\<langle>PK 0, pubkeysset (A 0)\<rangle>\<rangle>,
\<^cancel>\<open>rule R^10_2\<close>
\<langle>\<two>, send\<langle>Fun (attack 1) []\<rangle>\<rangle>
]"
theorem "stm.tfr\<^sub>s\<^sub>s\<^sub>t (unlabel S'\<^sub>k\<^sub>s)"
proof -
let ?S = "unlabel S'\<^sub>k\<^sub>s"
let ?M = "concat (map subterms_list (trms_list\<^sub>s\<^sub>s\<^sub>t ?S@map (pair' tuple) (setops_list\<^sub>s\<^sub>s\<^sub>t ?S)))"
have "comp_tfr\<^sub>s\<^sub>s\<^sub>t arity Ana \<Gamma> tuple ?M ?S" by eval
thus ?thesis by (rule stm.tfr\<^sub>s\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>s\<^sub>t)
qed
subsection \<open>Theorem: The steps of the keyserver protocols from the ESORICS18 paper satisfy the conditions for parallel composition\<close>
theorem
fixes S f
defines "S \<equiv> [PK 0, invkeymsg (PK 0), Fun encodingsecret []]@concat (
map (\<lambda>s. [s, Fun tuple [PK 0, s]])
[validset (A 0) (A 1), beginauthset 0 (A 0) (A 1), endauthset 0 (A 0) (A 1),
beginauthset 1 (A 0) (A 1), endauthset 1 (A 0) (A 1)])@
[A 0]"
and "f \<equiv> \<lambda>M. {t \<cdot> \<delta> | t \<delta>. t \<in> M \<and> tm.wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \<delta> \<and> im.wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \<delta>) \<and> fv (t \<cdot> \<delta>) = {}}"
and "Sec \<equiv> (f (set S)) - {m. im.intruder_synth {} m}"
shows "stm.par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t S'\<^sub>k\<^sub>s Sec"
proof -
let ?N = "\<lambda>P. concat (map subterms_list (trms_list\<^sub>s\<^sub>s\<^sub>t P@map (pair' tuple) (setops_list\<^sub>s\<^sub>s\<^sub>t P)))"
let ?M = "\<lambda>l. ?N (proj_unl l S'\<^sub>k\<^sub>s)"
have "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \<Gamma> tuple S'\<^sub>k\<^sub>s ?M S"
unfolding S_def by eval
thus ?thesis
using stm.par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t[of S'\<^sub>k\<^sub>s ?M S]
unfolding Sec_def f_def wf\<^sub>t\<^sub>r\<^sub>m_def[symmetric] by blast
qed
end

View File

@ -0,0 +1,305 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2015-2020
All Rights Reserved.
Redistribution and use in source and binary forms, with or without
modification, are permitted provided that the following conditions are
met:
- Redistributions of source code must retain the above copyright
notice, this list of conditions and the following disclaimer.
- Redistributions in binary form must reproduce the above copyright
notice, this list of conditions and the following disclaimer in the
documentation and/or other materials provided with the distribution.
- Neither the name of the copyright holder nor the names of its
contributors may be used to endorse or promote products
derived from this software without specific prior written
permission.
THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
"AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*)
(* Title: Example_TLS.thy
Author: Andreas Viktor Hess, DTU
*)
section \<open>Proving Type-Flaw Resistance of the TLS Handshake Protocol\<close>
theory Example_TLS
imports "../Typed_Model"
begin
declare [[code_timing]]
subsection \<open>TLS example: Datatypes and functions setup\<close>
datatype ex_atom = PrivKey | SymKey | PubConst | Agent | Nonce | Bot
datatype ex_fun =
clientHello | clientKeyExchange | clientFinished
| serverHello | serverCert | serverHelloDone
| finished | changeCipher | x509 | prfun | master | pmsForm
| sign | hash | crypt | pub | concat | privkey nat
| pubconst ex_atom nat
type_synonym ex_type = "(ex_fun, ex_atom) term_type"
type_synonym ex_var = "ex_type \<times> nat"
instance ex_atom::finite
proof
let ?S = "UNIV::ex_atom set"
have "?S = {PrivKey, SymKey, PubConst, Agent, Nonce, Bot}" by (auto intro: ex_atom.exhaust)
thus "finite ?S" by (metis finite.emptyI finite.insertI)
qed
type_synonym ex_term = "(ex_fun, ex_var) term"
type_synonym ex_terms = "(ex_fun, ex_var) terms"
primrec arity::"ex_fun \<Rightarrow> nat" where
"arity changeCipher = 0"
| "arity clientFinished = 4"
| "arity clientHello = 5"
| "arity clientKeyExchange = 1"
| "arity concat = 5"
| "arity crypt = 2"
| "arity finished = 1"
| "arity hash = 1"
| "arity master = 3"
| "arity pmsForm = 1"
| "arity prfun = 1"
| "arity (privkey _) = 0"
| "arity pub = 1"
| "arity (pubconst _ _) = 0"
| "arity serverCert = 1"
| "arity serverHello = 5"
| "arity serverHelloDone = 0"
| "arity sign = 2"
| "arity x509 = 2"
fun public::"ex_fun \<Rightarrow> bool" where
"public (privkey _) = False"
| "public _ = True"
fun Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t::"ex_term list \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t [Fun pub [k],m] = ([k], [m])"
| "Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t _ = ([], [])"
fun Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n::"ex_term list \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n [k,m] = ([], [m])"
| "Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n _ = ([], [])"
fun Ana::"ex_term \<Rightarrow> (ex_term list \<times> ex_term list)" where
"Ana (Fun crypt T) = Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t T"
| "Ana (Fun finished T) = ([], T)"
| "Ana (Fun master T) = ([], T)"
| "Ana (Fun pmsForm T) = ([], T)"
| "Ana (Fun serverCert T) = ([], T)"
| "Ana (Fun serverHello T) = ([], T)"
| "Ana (Fun sign T) = Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n T"
| "Ana (Fun x509 T) = ([], T)"
| "Ana _ = ([], [])"
subsection \<open>TLS example: Locale interpretation\<close>
lemma assm1:
"Ana t = (K,M) \<Longrightarrow> fv\<^sub>s\<^sub>e\<^sub>t (set K) \<subseteq> fv t"
"Ana t = (K,M) \<Longrightarrow> (\<And>g S'. Fun g S' \<sqsubseteq> t \<Longrightarrow> length S' = arity g)
\<Longrightarrow> k \<in> set K \<Longrightarrow> Fun f T' \<sqsubseteq> k \<Longrightarrow> length T' = arity f"
"Ana t = (K,M) \<Longrightarrow> K \<noteq> [] \<or> M \<noteq> [] \<Longrightarrow> Ana (t \<cdot> \<delta>) = (K \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>, M \<cdot>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<delta>)"
by (rule Ana.cases[of "t"], auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)+
lemma assm2: "Ana (Fun f T) = (K, M) \<Longrightarrow> set M \<subseteq> set T"
by (rule Ana.cases[of "Fun f T"]) (auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)
lemma assm6: "0 < arity f \<Longrightarrow> public f" by (cases f) simp_all
global_interpretation im: intruder_model arity public Ana
defines wf\<^sub>t\<^sub>r\<^sub>m = "im.wf\<^sub>t\<^sub>r\<^sub>m"
and wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s = "im.wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s"
by unfold_locales (metis assm1(1), metis assm1(2), rule Ana.simps, metis assm2, metis assm1(3))
subsection \<open>TLS Example: Typing function\<close>
definition \<Gamma>\<^sub>v::"ex_var \<Rightarrow> ex_type" where
"\<Gamma>\<^sub>v v = (if (\<forall>t \<in> subterms (fst v). case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True)
then fst v else TAtom Bot)"
fun \<Gamma>::"ex_term \<Rightarrow> ex_type" where
"\<Gamma> (Var v) = \<Gamma>\<^sub>v v"
| "\<Gamma> (Fun (privkey _) _) = TAtom PrivKey"
| "\<Gamma> (Fun changeCipher _) = TAtom PubConst"
| "\<Gamma> (Fun serverHelloDone _) = TAtom PubConst"
| "\<Gamma> (Fun (pubconst \<tau> _) _) = TAtom \<tau>"
| "\<Gamma> (Fun f T) = TComp f (map \<Gamma> T)"
subsection \<open>TLS Example: Locale interpretation (typed model)\<close>
lemma assm7: "arity c = 0 \<Longrightarrow> \<exists>a. \<forall>X. \<Gamma> (Fun c X) = TAtom a" by (cases c) simp_all
lemma assm8: "0 < arity f \<Longrightarrow> \<Gamma> (Fun f X) = TComp f (map \<Gamma> X)" by (cases f) simp_all
lemma assm9: "infinite {c. \<Gamma> (Fun c []) = TAtom a \<and> public c}"
proof -
let ?T = "(range (pubconst a))::ex_fun set"
have *:
"\<And>x y::nat. x \<in> UNIV \<Longrightarrow> y \<in> UNIV \<Longrightarrow> (pubconst a x = pubconst a y) = (x = y)"
"\<And>x::nat. x \<in> UNIV \<Longrightarrow> pubconst a x \<in> ?T"
"\<And>y::ex_fun. y \<in> ?T \<Longrightarrow> \<exists>x \<in> UNIV. y = pubconst a x"
by auto
have "?T \<subseteq> {c. \<Gamma> (Fun c []) = TAtom a \<and> public c}" by auto
moreover have "\<exists>f::nat \<Rightarrow> ex_fun. bij_betw f UNIV ?T"
using bij_betwI'[OF *] by blast
hence "infinite ?T" by (metis nat_not_finite bij_betw_finite)
ultimately show ?thesis using infinite_super by blast
qed
lemma assm10: "TComp f T \<sqsubseteq> \<Gamma> t \<Longrightarrow> arity f > 0"
proof (induction rule: \<Gamma>.induct)
case (1 x)
hence *: "TComp f T \<sqsubseteq> \<Gamma>\<^sub>v x" by simp
hence "\<Gamma>\<^sub>v x \<noteq> TAtom Bot" unfolding \<Gamma>\<^sub>v_def by force
hence "\<forall>t \<in> subterms (fst x). case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True"
unfolding \<Gamma>\<^sub>v_def by argo
thus ?case using * unfolding \<Gamma>\<^sub>v_def by fastforce
qed auto
lemma assm11: "im.wf\<^sub>t\<^sub>r\<^sub>m (\<Gamma> (Var x))"
proof -
have "im.wf\<^sub>t\<^sub>r\<^sub>m (\<Gamma>\<^sub>v x)" unfolding \<Gamma>\<^sub>v_def im.wf\<^sub>t\<^sub>r\<^sub>m_def by auto
thus ?thesis by simp
qed
lemma assm12: "\<Gamma> (Var (\<tau>, n)) = \<Gamma> (Var (\<tau>, m))"
apply (cases "\<forall>t \<in> subterms \<tau>. case t of
(TComp f T) \<Rightarrow> arity f > 0 \<and> arity f = length T
| _ \<Rightarrow> True")
by (auto simp add: \<Gamma>\<^sub>v_def)
lemma Ana_const: "arity c = 0 \<Longrightarrow> Ana (Fun c T) = ([],[])"
by (cases c) simp_all
lemma Ana_keys_subterm: "Ana t = (K,T) \<Longrightarrow> k \<in> set K \<Longrightarrow> k \<sqsubset> t"
proof (induct t rule: Ana.induct)
case (1 U)
then obtain m where "U = [Fun pub [k], m]" "K = [k]" "T = [m]"
by (auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)
thus ?case using Fun_subterm_inside_params[of k crypt U] by auto
qed (auto elim!: Ana\<^sub>c\<^sub>r\<^sub>y\<^sub>p\<^sub>t.elims Ana\<^sub>s\<^sub>i\<^sub>g\<^sub>n.elims)
global_interpretation tm: typed_model' arity public Ana \<Gamma>
by (unfold_locales, unfold wf\<^sub>t\<^sub>r\<^sub>m_def[symmetric],
metis assm7, metis assm8, metis assm9, metis assm10, metis assm11, metis assm6,
metis assm12, metis Ana_const, metis Ana_keys_subterm)
subsection \<open>TLS example: Proving type-flaw resistance\<close>
abbreviation \<Gamma>\<^sub>v_clientHello where
"\<Gamma>\<^sub>v_clientHello \<equiv>
TComp clientHello [TAtom Nonce, TAtom Nonce, TAtom Nonce, TAtom Nonce, TAtom Nonce]"
abbreviation \<Gamma>\<^sub>v_serverHello where
"\<Gamma>\<^sub>v_serverHello \<equiv>
TComp serverHello [TAtom Nonce, TAtom Nonce, TAtom Nonce, TAtom Nonce, TAtom Nonce]"
abbreviation \<Gamma>\<^sub>v_pub where
"\<Gamma>\<^sub>v_pub \<equiv> TComp pub [TAtom PrivKey]"
abbreviation \<Gamma>\<^sub>v_x509 where
"\<Gamma>\<^sub>v_x509 \<equiv> TComp x509 [TAtom Agent, \<Gamma>\<^sub>v_pub]"
abbreviation \<Gamma>\<^sub>v_sign where
"\<Gamma>\<^sub>v_sign \<equiv> TComp sign [TAtom PrivKey, \<Gamma>\<^sub>v_x509]"
abbreviation \<Gamma>\<^sub>v_serverCert where
"\<Gamma>\<^sub>v_serverCert \<equiv> TComp serverCert [\<Gamma>\<^sub>v_sign]"
abbreviation \<Gamma>\<^sub>v_pmsForm where
"\<Gamma>\<^sub>v_pmsForm \<equiv> TComp pmsForm [TAtom SymKey]"
abbreviation \<Gamma>\<^sub>v_crypt where
"\<Gamma>\<^sub>v_crypt \<equiv> TComp crypt [\<Gamma>\<^sub>v_pub, \<Gamma>\<^sub>v_pmsForm]"
abbreviation \<Gamma>\<^sub>v_clientKeyExchange where
"\<Gamma>\<^sub>v_clientKeyExchange \<equiv>
TComp clientKeyExchange [\<Gamma>\<^sub>v_crypt]"
abbreviation \<Gamma>\<^sub>v_HSMsgs where
"\<Gamma>\<^sub>v_HSMsgs \<equiv> TComp concat [
\<Gamma>\<^sub>v_clientHello,
\<Gamma>\<^sub>v_serverHello,
\<Gamma>\<^sub>v_serverCert,
TAtom PubConst,
\<Gamma>\<^sub>v_clientKeyExchange]"
(* Variables from TLS *)
abbreviation "T\<^sub>1 n \<equiv> Var (TAtom Nonce,n)"
abbreviation "T\<^sub>2 n \<equiv> Var (TAtom Nonce,n)"
abbreviation "R\<^sub>A n \<equiv> Var (TAtom Nonce,n)"
abbreviation "R\<^sub>B n \<equiv> Var (TAtom Nonce,n)"
abbreviation "S n \<equiv> Var (TAtom Nonce,n)"
abbreviation "Cipher n \<equiv> Var (TAtom Nonce,n)"
abbreviation "Comp n \<equiv> Var (TAtom Nonce,n)"
abbreviation "B n \<equiv> Var (TAtom Agent,n)"
abbreviation "Pr\<^sub>c\<^sub>a n \<equiv> Var (TAtom PrivKey,n)"
abbreviation "PMS n \<equiv> Var (TAtom SymKey,n)"
abbreviation "P\<^sub>B n \<equiv> Var (TComp pub [TAtom PrivKey],n)"
abbreviation "HSMsgs n \<equiv> Var (\<Gamma>\<^sub>v_HSMsgs,n)"
subsubsection \<open>Defining the over-approximation set\<close>
abbreviation clientHello\<^sub>t\<^sub>r\<^sub>m where
"clientHello\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun clientHello [T\<^sub>1 0, R\<^sub>A 1, S 2, Cipher 3, Comp 4]"
abbreviation serverHello\<^sub>t\<^sub>r\<^sub>m where
"serverHello\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun serverHello [T\<^sub>2 0, R\<^sub>B 1, S 2, Cipher 3, Comp 4]"
abbreviation serverCert\<^sub>t\<^sub>r\<^sub>m where
"serverCert\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun serverCert [Fun sign [Pr\<^sub>c\<^sub>a 0, Fun x509 [B 1, P\<^sub>B 2]]]"
abbreviation serverHelloDone\<^sub>t\<^sub>r\<^sub>m where
"serverHelloDone\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun serverHelloDone []"
abbreviation clientKeyExchange\<^sub>t\<^sub>r\<^sub>m where
"clientKeyExchange\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun clientKeyExchange [Fun crypt [P\<^sub>B 0, Fun pmsForm [PMS 1]]]"
abbreviation changeCipher\<^sub>t\<^sub>r\<^sub>m where
"changeCipher\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun changeCipher []"
abbreviation finished\<^sub>t\<^sub>r\<^sub>m where
"finished\<^sub>t\<^sub>r\<^sub>m \<equiv> Fun finished [Fun prfun [
Fun clientFinished [
Fun prfun [Fun master [PMS 0, R\<^sub>A 1, R\<^sub>B 2]],
R\<^sub>A 3, R\<^sub>B 4, Fun hash [HSMsgs 5]
]
]]"
definition M\<^sub>T\<^sub>L\<^sub>S::"ex_term list" where
"M\<^sub>T\<^sub>L\<^sub>S \<equiv> [
clientHello\<^sub>t\<^sub>r\<^sub>m,
serverHello\<^sub>t\<^sub>r\<^sub>m,
serverCert\<^sub>t\<^sub>r\<^sub>m,
serverHelloDone\<^sub>t\<^sub>r\<^sub>m,
clientKeyExchange\<^sub>t\<^sub>r\<^sub>m,
changeCipher\<^sub>t\<^sub>r\<^sub>m,
finished\<^sub>t\<^sub>r\<^sub>m
]"
subsection \<open>Theorem: The TLS handshake protocol is type-flaw resistant\<close>
theorem "tm.tfr\<^sub>s\<^sub>e\<^sub>t (set M\<^sub>T\<^sub>L\<^sub>S)"
by (rule tm.tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t') eval
end