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

This commit is contained in:
Achim D. Brucker 2020-05-23 15:44:53 +01:00
commit 216e3831c7
47 changed files with 24322 additions and 0 deletions

View File

@ -0,0 +1,110 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Eisbach_Protocol_Verification.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section \<open>Useful Eisbach Methods for Automating Protocol Verification\<close>
theory Eisbach_Protocol_Verification
imports Main "HOL-Eisbach.Eisbach_Tools"
begin
named_theorems exhausts
named_theorems type_class_instance_lemmata
named_theorems protocol_checks
named_theorems coverage_check_unfold_protocol_lemma
named_theorems coverage_check_unfold_lemmata
named_theorems coverage_check_intro_lemmata
named_theorems transaction_coverage_lemmata
method UNIV_lemma =
(rule UNIV_eq_I; (subst insert_iff)+; subst empty_iff; smt exhausts)+
method type_class_instance =
(intro_classes; auto simp add: type_class_instance_lemmata)
method protocol_model_subgoal =
(((rule allI, case_tac f); (erule forw_subst)+)?; simp_all)
method protocol_model_interpretation =
(unfold_locales; protocol_model_subgoal+)
method check_protocol_intro =
(unfold_locales, unfold protocol_checks[symmetric])
method check_protocol_with methods meth =
(check_protocol_intro, meth)
method check_protocol' =
(check_protocol_with \<open>code_simp+\<close>)
method check_protocol_unsafe' =
(check_protocol_with \<open>eval+\<close>)
method check_protocol =
(check_protocol_with \<open>
code_simp,
code_simp,
code_simp,
code_simp,
code_simp\<close>)
method check_protocol_unsafe =
(check_protocol_with \<open>
eval,
eval,
eval,
eval,
eval\<close>)
method coverage_check_intro =
(((unfold coverage_check_unfold_protocol_lemma)?;
intro coverage_check_intro_lemmata;
simp only: list_all_simps list_all_append list.map concat.simps map_append product_concat_map;
intro conjI TrueI);
(clarsimp+)?;
((rule transaction_coverage_lemmata)+)?)
method coverage_check_unfold =
(unfold coverage_check_unfold_protocol_lemma coverage_check_unfold_lemmata
list_all_iff Let_def case_prod_unfold Product_Type.fst_conv Product_Type.snd_conv)
end

View File

@ -0,0 +1,54 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Examples.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>Examples\<close>
theory Examples
imports "examples/Keyserver"
"examples/Keyserver2"
"examples/Keyserver_Composition"
"examples/PKCS/PKCS_Model03"
"examples/PKCS/PKCS_Model07"
"examples/PKCS/PKCS_Model09"
begin
end

View File

@ -0,0 +1,53 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: PSPSP.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>PSPSP\<close>
theory PSPSP
imports "Stateful_Protocol_Verification"
"Eisbach_Protocol_Verification"
"trac/trac"
begin
end

View File

@ -0,0 +1,16 @@
chapter AFP
session "Automated_Stateful_Protocol_Verification-devel" (AFP) = "Stateful_Protocol_Composition_and_Typing" +
options [timeout = 2400]
sessions
"HOL-Eisbach"
directories
"trac"
"examples"
"examples/PKCS"
theories
"PSPSP"
"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

View File

@ -0,0 +1,246 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Term_Abstraction.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>Term Abstraction\<close>
theory Term_Abstraction
imports Transactions
begin
subsection \<open>Definitions\<close>
fun to_abs ("\<alpha>\<^sub>0") where
"\<alpha>\<^sub>0 [] _ = {}"
| "\<alpha>\<^sub>0 ((Fun (Val m) [],Fun (Set s) S)#D) n =
(if m = n then insert s (\<alpha>\<^sub>0 D n) else \<alpha>\<^sub>0 D n)"
| "\<alpha>\<^sub>0 (_#D) n = \<alpha>\<^sub>0 D n"
fun abs_apply_term (infixl "\<cdot>\<^sub>\<alpha>" 67) where
"Var x \<cdot>\<^sub>\<alpha> \<alpha> = Var x"
| "Fun (Val n) T \<cdot>\<^sub>\<alpha> \<alpha> = Fun (Abs (\<alpha> n)) (map (\<lambda>t. t \<cdot>\<^sub>\<alpha> \<alpha>) T)"
| "Fun f T \<cdot>\<^sub>\<alpha> \<alpha> = Fun f (map (\<lambda>t. t \<cdot>\<^sub>\<alpha> \<alpha>) T)"
definition abs_apply_list (infixl "\<cdot>\<^sub>\<alpha>\<^sub>l\<^sub>i\<^sub>s\<^sub>t" 67) where
"M \<cdot>\<^sub>\<alpha>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<alpha> \<equiv> map (\<lambda>t. t \<cdot>\<^sub>\<alpha> \<alpha>) M"
definition abs_apply_terms (infixl "\<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t" 67) where
"M \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t \<alpha> \<equiv> (\<lambda>t. t \<cdot>\<^sub>\<alpha> \<alpha>) ` M"
definition abs_apply_pairs (infixl "\<cdot>\<^sub>\<alpha>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s" 67) where
"F \<cdot>\<^sub>\<alpha>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \<alpha> \<equiv> map (\<lambda>(s,t). (s \<cdot>\<^sub>\<alpha> \<alpha>, t \<cdot>\<^sub>\<alpha> \<alpha>)) F"
definition abs_apply_strand_step (infixl "\<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>t\<^sub>p" 67) where
"s \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>t\<^sub>p \<alpha> \<equiv> (case s of
(l,send\<langle>t\<rangle>) \<Rightarrow> (l,send\<langle>t \<cdot>\<^sub>\<alpha> \<alpha>\<rangle>)
| (l,receive\<langle>t\<rangle>) \<Rightarrow> (l,receive\<langle>t \<cdot>\<^sub>\<alpha> \<alpha>\<rangle>)
| (l,\<langle>ac: t \<doteq> t'\<rangle>) \<Rightarrow> (l,\<langle>ac: (t \<cdot>\<^sub>\<alpha> \<alpha>) \<doteq> (t' \<cdot>\<^sub>\<alpha> \<alpha>)\<rangle>)
| (l,insert\<langle>t,t'\<rangle>) \<Rightarrow> (l,insert\<langle>t \<cdot>\<^sub>\<alpha> \<alpha>,t' \<cdot>\<^sub>\<alpha> \<alpha>\<rangle>)
| (l,delete\<langle>t,t'\<rangle>) \<Rightarrow> (l,delete\<langle>t \<cdot>\<^sub>\<alpha> \<alpha>,t' \<cdot>\<^sub>\<alpha> \<alpha>\<rangle>)
| (l,\<langle>ac: t \<in> t'\<rangle>) \<Rightarrow> (l,\<langle>ac: (t \<cdot>\<^sub>\<alpha> \<alpha>) \<in> (t' \<cdot>\<^sub>\<alpha> \<alpha>)\<rangle>)
| (l,\<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: F'\<rangle>) \<Rightarrow> (l,\<forall>X\<langle>\<or>\<noteq>: (F \<cdot>\<^sub>\<alpha>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \<alpha>) \<or>\<notin>: (F' \<cdot>\<^sub>\<alpha>\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \<alpha>)\<rangle>))"
definition abs_apply_strand (infixl "\<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>t" 67) where
"S \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>t \<alpha> \<equiv> map (\<lambda>x. x \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>t\<^sub>p \<alpha>) S"
subsection \<open>Lemmata\<close>
lemma to_abs_alt_def:
"\<alpha>\<^sub>0 D n = {s. \<exists>S. (Fun (Val n) [], Fun (Set s) S) \<in> set D}"
by (induct D n rule: to_abs.induct) auto
lemma abs_term_apply_const[simp]:
"is_Val f \<Longrightarrow> Fun f [] \<cdot>\<^sub>\<alpha> a = Fun (Abs (a (the_Val f))) []"
"\<not>is_Val f \<Longrightarrow> Fun f [] \<cdot>\<^sub>\<alpha> a = Fun f []"
by (cases f; auto)+
lemma abs_fv: "fv (t \<cdot>\<^sub>\<alpha> a) = fv t"
by (induct t a rule: abs_apply_term.induct) auto
lemma abs_eq_if_no_Val:
assumes "\<forall>f \<in> funs_term t. \<not>is_Val f"
shows "t \<cdot>\<^sub>\<alpha> a = t \<cdot>\<^sub>\<alpha> b"
using assms
proof (induction t)
case (Fun f T) thus ?case by (cases f) simp_all
qed simp
lemma abs_list_set_is_set_abs_set: "set (M \<cdot>\<^sub>\<alpha>\<^sub>l\<^sub>i\<^sub>s\<^sub>t \<alpha>) = (set M) \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t \<alpha>"
unfolding abs_apply_list_def abs_apply_terms_def by simp
lemma abs_set_empty[simp]: "{} \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t \<alpha> = {}"
unfolding abs_apply_terms_def by simp
lemma abs_in:
assumes "t \<in> M"
shows "t \<cdot>\<^sub>\<alpha> \<alpha> \<in> M \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t \<alpha>"
using assms unfolding abs_apply_terms_def
by (induct t \<alpha> rule: abs_apply_term.induct) blast+
lemma abs_set_union: "(A \<union> B) \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a = (A \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a) \<union> (B \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a)"
unfolding abs_apply_terms_def
by auto
lemma abs_subterms: "subterms (t \<cdot>\<^sub>\<alpha> \<alpha>) = subterms t \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t \<alpha>"
proof (induction t)
case (Fun f T) thus ?case by (cases f) (auto simp add: abs_apply_terms_def)
qed (simp add: abs_apply_terms_def)
lemma abs_subterms_in: "s \<in> subterms t \<Longrightarrow> s \<cdot>\<^sub>\<alpha> a \<in> subterms (t \<cdot>\<^sub>\<alpha> a)"
proof (induction t)
case (Fun f T) thus ?case by (cases f) auto
qed simp
lemma abs_ik_append: "(ik\<^sub>s\<^sub>s\<^sub>t (A@B) \<cdot>\<^sub>s\<^sub>e\<^sub>t I) \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a = (ik\<^sub>s\<^sub>s\<^sub>t A \<cdot>\<^sub>s\<^sub>e\<^sub>t I) \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a \<union> (ik\<^sub>s\<^sub>s\<^sub>t B \<cdot>\<^sub>s\<^sub>e\<^sub>t I) \<cdot>\<^sub>\<alpha>\<^sub>s\<^sub>e\<^sub>t a"
unfolding abs_apply_terms_def ik\<^sub>s\<^sub>s\<^sub>t_def
by auto
lemma to_abs_in:
assumes "(Fun (Val n) [], Fun (Set s) []) \<in> set D"
shows "s \<in> \<alpha>\<^sub>0 D n"
using assms by (induct rule: to_abs.induct) auto
lemma to_abs_empty_iff_notin_db:
"Fun (Val n) [] \<cdot>\<^sub>\<alpha> \<alpha>\<^sub>0 D = Fun (Abs {}) [] \<longleftrightarrow> (\<nexists>s S. (Fun (Val n) [], Fun (Set s) S) \<in> set D)"
by (simp add: to_abs_alt_def)
lemma to_abs_list_insert:
assumes "Fun (Val n) [] \<noteq> t"
shows "\<alpha>\<^sub>0 D n = \<alpha>\<^sub>0 (List.insert (t,s) D) n"
using assms to_abs_alt_def[of D n] to_abs_alt_def[of "List.insert (t,s) D" n]
by auto
lemma to_abs_list_insert':
"insert s (\<alpha>\<^sub>0 D n) = \<alpha>\<^sub>0 (List.insert (Fun (Val n) [], Fun (Set s) S) D) n"
using to_abs_alt_def[of D n]
to_abs_alt_def[of "List.insert (Fun (Val n) [], Fun (Set s) S) D" n]
by auto
lemma to_abs_list_remove_all:
assumes "Fun (Val n) [] \<noteq> t"
shows "\<alpha>\<^sub>0 D n = \<alpha>\<^sub>0 (List.removeAll (t,s) D) n"
using assms to_abs_alt_def[of D n] to_abs_alt_def[of "List.removeAll (t,s) D" n]
by auto
lemma to_abs_list_remove_all':
"\<alpha>\<^sub>0 D n - {s} = \<alpha>\<^sub>0 (filter (\<lambda>d. \<nexists>S. d = (Fun (Val n) [], Fun (Set s) S)) D) n"
using to_abs_alt_def[of D n]
to_abs_alt_def[of "filter (\<lambda>d. \<nexists>S. d = (Fun (Val n) [], Fun (Set s) S)) D" n]
by auto
lemma to_abs_db\<^sub>s\<^sub>s\<^sub>t_append:
assumes "\<forall>u s. insert\<langle>u, s\<rangle> \<in> set B \<longrightarrow> Fun (Val n) [] \<noteq> u \<cdot> \<I>"
and "\<forall>u s. delete\<langle>u, s\<rangle> \<in> set B \<longrightarrow> Fun (Val n) [] \<noteq> u \<cdot> \<I>"
shows "\<alpha>\<^sub>0 (db'\<^sub>s\<^sub>s\<^sub>t A \<I> D) n = \<alpha>\<^sub>0 (db'\<^sub>s\<^sub>s\<^sub>t (A@B) \<I> D) n"
using assms
proof (induction B rule: List.rev_induct)
case (snoc b B)
hence IH: "\<alpha>\<^sub>0 (db'\<^sub>s\<^sub>s\<^sub>t A \<I> D) n = \<alpha>\<^sub>0 (db'\<^sub>s\<^sub>s\<^sub>t (A@B) \<I> D) n" by auto
have *: "\<forall>u s. b = insert\<langle>u,s\<rangle> \<longrightarrow> Fun (Val n) [] \<noteq> u \<cdot> \<I>"
"\<forall>u s. b = delete\<langle>u,s\<rangle> \<longrightarrow> Fun (Val n) [] \<noteq> u \<cdot> \<I>"
using snoc.prems by simp_all
show ?case
proof (cases b)
case (Insert u s)
hence **: "db'\<^sub>s\<^sub>s\<^sub>t (A@B@[b]) \<I> D = List.insert (u \<cdot> \<I>,s \<cdot> \<I>) (db'\<^sub>s\<^sub>s\<^sub>t (A@B) \<I> D)"
using db\<^sub>s\<^sub>s\<^sub>t_append[of "A@B" "[b]"] by simp
have "Fun (Val n) [] \<noteq> u \<cdot> \<I>" using *(1) Insert by auto
thus ?thesis using IH ** to_abs_list_insert by metis
next
case (Delete u s)
hence **: "db'\<^sub>s\<^sub>s\<^sub>t (A@B@[b]) \<I> D = List.removeAll (u \<cdot> \<I>,s \<cdot> \<I>) (db'\<^sub>s\<^sub>s\<^sub>t (A@B) \<I> D)"
using db\<^sub>s\<^sub>s\<^sub>t_append[of "A@B" "[b]"] by simp
have "Fun (Val n) [] \<noteq> u \<cdot> \<I>" using *(2) Delete by auto
thus ?thesis using IH ** to_abs_list_remove_all by metis
qed (simp_all add: db\<^sub>s\<^sub>s\<^sub>t_no_upd_append[of "[b]" "A@B"] IH)
qed simp
lemma to_abs_neq_imp_db_update:
assumes "\<alpha>\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t A I) n \<noteq> \<alpha>\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t (A@B) I) n"
shows "\<exists>u s. u \<cdot> I = Fun (Val n) [] \<and> (insert\<langle>u,s\<rangle> \<in> set B \<or> delete\<langle>u,s\<rangle> \<in> set B)"
proof -
{ fix D have ?thesis when "\<alpha>\<^sub>0 D n \<noteq> \<alpha>\<^sub>0 (db'\<^sub>s\<^sub>s\<^sub>t B I D) n" using that
proof (induction B I D rule: db'\<^sub>s\<^sub>s\<^sub>t.induct)
case 2 thus ?case
by (metis db'\<^sub>s\<^sub>s\<^sub>t.simps(2) list.set_intros(1,2) subst_apply_pair_pair to_abs_list_insert)
next
case 3 thus ?case
by (metis db'\<^sub>s\<^sub>s\<^sub>t.simps(3) list.set_intros(1,2) subst_apply_pair_pair to_abs_list_remove_all)
qed simp_all
} thus ?thesis using assms by (metis db\<^sub>s\<^sub>s\<^sub>t_append db\<^sub>s\<^sub>s\<^sub>t_def)
qed
lemma abs_term_subst_eq:
fixes \<delta> \<theta>::"(('a,'b,'c) prot_fun, ('d,'e prot_atom) term \<times> nat) subst"
assumes "\<forall>x \<in> fv t. \<delta> x \<cdot>\<^sub>\<alpha> a = \<theta> x \<cdot>\<^sub>\<alpha> b"
and "\<nexists>n T. Fun (Val n) T \<in> subterms t"
shows "t \<cdot> \<delta> \<cdot>\<^sub>\<alpha> a = t \<cdot> \<theta> \<cdot>\<^sub>\<alpha> b"
using assms
proof (induction t)
case (Fun f T) thus ?case
proof (cases f)
case (Val n)
hence False using Fun.prems(2) by blast
thus ?thesis by metis
qed auto
qed simp
lemma abs_term_subst_eq':
fixes \<delta> \<theta>::"(('a,'b,'c) prot_fun, ('d,'e prot_atom) term \<times> nat) subst"
assumes "\<forall>x \<in> fv t. \<delta> x \<cdot>\<^sub>\<alpha> a = \<theta> x"
and "\<nexists>n T. Fun (Val n) T \<in> subterms t"
shows "t \<cdot> \<delta> \<cdot>\<^sub>\<alpha> a = t \<cdot> \<theta>"
using assms
proof (induction t)
case (Fun f T) thus ?case
proof (cases f)
case (Val n)
hence False using Fun.prems(2) by blast
thus ?thesis by metis
qed auto
qed simp
lemma abs_val_in_funs_term:
assumes "f \<in> funs_term t" "is_Val f"
shows "Abs (\<alpha> (the_Val f)) \<in> funs_term (t \<cdot>\<^sub>\<alpha> \<alpha>)"
using assms by (induct t \<alpha> rule: abs_apply_term.induct) auto
end

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,451 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Term_Variants.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>Term Variants\<close>
theory Term_Variants
imports Stateful_Protocol_Composition_and_Typing.Intruder_Deduction
begin
fun term_variants where
"term_variants P (Var x) = [Var x]"
| "term_variants P (Fun f T) = (
let S = product_lists (map (term_variants P) T)
in map (Fun f) S@concat (map (\<lambda>g. map (Fun g) S) (P f)))"
inductive term_variants_pred where
term_variants_Var:
"term_variants_pred P (Var x) (Var x)"
| term_variants_P:
"\<lbrakk>length T = length S; \<And>i. i < length T \<Longrightarrow> term_variants_pred P (T ! i) (S ! i); g \<in> set (P f)\<rbrakk>
\<Longrightarrow> term_variants_pred P (Fun f T) (Fun g S)"
| term_variants_Fun:
"\<lbrakk>length T = length S; \<And>i. i < length T \<Longrightarrow> term_variants_pred P (T ! i) (S ! i)\<rbrakk>
\<Longrightarrow> term_variants_pred P (Fun f T) (Fun f S)"
lemma term_variants_pred_inv:
assumes "term_variants_pred P (Fun f T) (Fun h S)"
shows "length T = length S"
and "\<And>i. i < length T \<Longrightarrow> term_variants_pred P (T ! i) (S ! i)"
and "f \<noteq> h \<Longrightarrow> h \<in> set (P f)"
using assms by (auto elim: term_variants_pred.cases)
lemma term_variants_pred_inv':
assumes "term_variants_pred P (Fun f T) t"
shows "is_Fun t"
and "length T = length (args t)"
and "\<And>i. i < length T \<Longrightarrow> term_variants_pred P (T ! i) (args t ! i)"
and "f \<noteq> the_Fun t \<Longrightarrow> the_Fun t \<in> set (P f)"
and "P \<equiv> (\<lambda>_. [])(g := [h]) \<Longrightarrow> f \<noteq> the_Fun t \<Longrightarrow> f = g \<and> the_Fun t = h"
using assms by (auto elim: term_variants_pred.cases)
lemma term_variants_pred_inv'':
assumes "term_variants_pred P t (Fun f T)"
shows "is_Fun t"
and "length T = length (args t)"
and "\<And>i. i < length T \<Longrightarrow> term_variants_pred P (args t ! i) (T ! i)"
and "f \<noteq> the_Fun t \<Longrightarrow> f \<in> set (P (the_Fun t))"
and "P \<equiv> (\<lambda>_. [])(g := [h]) \<Longrightarrow> f \<noteq> the_Fun t \<Longrightarrow> f = h \<and> the_Fun t = g"
using assms by (auto elim: term_variants_pred.cases)
lemma term_variants_pred_inv_Var:
"term_variants_pred P (Var x) t \<longleftrightarrow> t = Var x"
"term_variants_pred P t (Var x) \<longleftrightarrow> t = Var x"
by (auto intro: term_variants_Var elim: term_variants_pred.cases)
lemma term_variants_pred_inv_const:
"term_variants_pred P (Fun c []) t \<longleftrightarrow> ((\<exists>g \<in> set (P c). t = Fun g []) \<or> (t = Fun c []))"
by (auto intro: term_variants_P term_variants_Fun elim: term_variants_pred.cases)
lemma term_variants_pred_refl: "term_variants_pred P t t"
by (induct t) (auto intro: term_variants_pred.intros)
lemma term_variants_pred_refl_inv:
assumes st: "term_variants_pred P s t"
and P: "\<forall>f. \<forall>g \<in> set (P f). f = g"
shows "s = t"
using st P
proof (induction s t rule: term_variants_pred.induct)
case (term_variants_Var P x) thus ?case by blast
next
case (term_variants_P T S P g f)
hence "T ! i = S ! i" when i: "i < length T" for i using i by blast
hence "T = S" using term_variants_P.hyps(1) by (simp add: nth_equalityI)
thus ?case using term_variants_P.prems term_variants_P.hyps(3) by fast
next
case (term_variants_Fun T S P f)
hence "T ! i = S ! i" when i: "i < length T" for i using i by blast
hence "T = S" using term_variants_Fun.hyps(1) by (simp add: nth_equalityI)
thus ?case by fast
qed
lemma term_variants_pred_const:
assumes "b \<in> set (P a)"
shows "term_variants_pred P (Fun a []) (Fun b [])"
using term_variants_P[of "[]" "[]"] assms by simp
lemma term_variants_pred_const_cases:
"P a \<noteq> [] \<Longrightarrow> term_variants_pred P (Fun a []) t \<longleftrightarrow>
(t = Fun a [] \<or> (\<exists>b \<in> set (P a). t = Fun b []))"
"P a = [] \<Longrightarrow> term_variants_pred P (Fun a []) t \<longleftrightarrow> t = Fun a []"
using term_variants_pred_inv_const[of P] by auto
lemma term_variants_pred_param:
assumes "term_variants_pred P t s"
and fg: "f = g \<or> g \<in> set (P f)"
shows "term_variants_pred P (Fun f (S@t#T)) (Fun g (S@s#T))"
proof -
have 1: "length (S@t#T) = length (S@s#T)" by simp
have "term_variants_pred P (T ! i) (T ! i)" "term_variants_pred P (S ! i) (S ! i)" for i
by (metis term_variants_pred_refl)+
hence 2: "term_variants_pred P ((S@t#T) ! i) ((S@s#T) ! i)" for i
by (simp add: assms nth_Cons' nth_append)
show ?thesis by (metis term_variants_Fun[OF 1 2] term_variants_P[OF 1 2] fg)
qed
lemma term_variants_pred_Cons:
assumes t: "term_variants_pred P t s"
and T: "term_variants_pred P (Fun f T) (Fun f S)"
and fg: "f = g \<or> g \<in> set (P f)"
shows "term_variants_pred P (Fun f (t#T)) (Fun g (s#S))"
proof -
have 1: "length (t#T) = length (s#S)"
and "\<And>i. i < length T \<Longrightarrow> term_variants_pred P (T ! i) (S ! i)"
using term_variants_pred_inv[OF T] by simp_all
hence 2: "\<And>i. i < length (t#T) \<Longrightarrow> term_variants_pred P ((t#T) ! i) ((s#S) ! i)"
by (metis t One_nat_def diff_less length_Cons less_Suc_eq less_imp_diff_less nth_Cons'
zero_less_Suc)
show ?thesis using 1 2 fg by (auto intro: term_variants_pred.intros)
qed
lemma term_variants_pred_dense:
fixes P Q::"'a set" and fs gs::"'a list"
defines "P_fs x \<equiv> if x \<in> P then fs else []"
and "P_gs x \<equiv> if x \<in> P then gs else []"
and "Q_fs x \<equiv> if x \<in> Q then fs else []"
assumes ut: "term_variants_pred P_fs u t"
and g: "g \<in> Q" "g \<in> set gs"
shows "\<exists>s. term_variants_pred P_gs u s \<and> term_variants_pred Q_fs s t"
proof -
define F where "F \<equiv> \<lambda>(P::'a set) (fs::'a list) x. if x \<in> P then fs else []"
show ?thesis using ut g P_fs_def unfolding P_gs_def Q_fs_def
proof (induction P_fs u t arbitrary: g gs rule: term_variants_pred.induct)
case (term_variants_Var P h x) thus ?case
by (auto intro: term_variants_pred.term_variants_Var)
next
case (term_variants_P T S P' h' h g gs)
note hyps = term_variants_P.hyps(1,2,4,5,6,7)
note IH = term_variants_P.hyps(3)
have "\<exists>s. term_variants_pred (F P gs) (T ! i) s \<and> term_variants_pred (F Q fs) s (S ! i)"
when i: "i < length T" for i
using IH[OF i hyps(4,5,6)] unfolding F_def by presburger
then obtain U where U:
"length T = length U" "\<And>i. i < length T \<Longrightarrow> term_variants_pred (F P gs) (T ! i) (U ! i)"
"length U = length S" "\<And>i. i < length U \<Longrightarrow> term_variants_pred (F Q fs) (U ! i) (S ! i)"
using hyps(1) Skolem_list_nth[of _ "\<lambda>i s. term_variants_pred (F P gs) (T ! i) s \<and>
term_variants_pred (F Q fs) s (S ! i)"]
by moura
show ?case
using term_variants_pred.term_variants_P[OF U(1,2), of g h]
term_variants_pred.term_variants_P[OF U(3,4), of h' g]
hyps(3)[unfolded hyps(6)] hyps(4,5)
unfolding F_def by force
next
case (term_variants_Fun T S P' h' g gs)
note hyps = term_variants_Fun.hyps(1,2,4,5,6)
note IH = term_variants_Fun.hyps(3)
have "\<exists>s. term_variants_pred (F P gs) (T ! i) s \<and> term_variants_pred (F Q fs) s (S ! i)"
when i: "i < length T" for i
using IH[OF i hyps(3,4,5)] unfolding F_def by presburger
then obtain U where U:
"length T = length U" "\<And>i. i < length T \<Longrightarrow> term_variants_pred (F P gs) (T ! i) (U ! i)"
"length U = length S" "\<And>i. i < length U \<Longrightarrow> term_variants_pred (F Q fs) (U ! i) (S ! i)"
using hyps(1) Skolem_list_nth[of _ "\<lambda>i s. term_variants_pred (F P gs) (T ! i) s \<and>
term_variants_pred (F Q fs) s (S ! i)"]
by moura
thus ?case
using term_variants_pred.term_variants_Fun[OF U(1,2)]
term_variants_pred.term_variants_Fun[OF U(3,4)]
unfolding F_def by meson
qed
qed
lemma term_variants_pred_dense':
assumes ut: "term_variants_pred ((\<lambda>_. [])(a := [b])) u t"
shows "\<exists>s. term_variants_pred ((\<lambda>_. [])(a := [c])) u s \<and>
term_variants_pred ((\<lambda>_. [])(c := [b])) s t"
using ut term_variants_pred_dense[of "{a}" "[b]" u t c "{c}" "[c]"]
unfolding fun_upd_def by simp
lemma term_variants_pred_eq_case:
fixes t s::"('a,'b) term"
assumes "term_variants_pred P t s" "\<forall>f \<in> funs_term t. P f = []"
shows "t = s"
using assms
proof (induction P t s rule: term_variants_pred.induct)
case (term_variants_Fun T S P f) thus ?case
using subtermeq_imp_funs_term_subset[OF Fun_param_in_subterms[OF nth_mem], of _ T f]
nth_equalityI[of T S]
by blast
qed (simp_all add: term_variants_pred_refl)
lemma term_variants_pred_subst:
assumes "term_variants_pred P t s"
shows "term_variants_pred P (t \<cdot> \<delta>) (s \<cdot> \<delta>)"
using assms
proof (induction P t s rule: term_variants_pred.induct)
case (term_variants_P T S P f g)
have 1: "length (map (\<lambda>t. t \<cdot> \<delta>) T) = length (map (\<lambda>t. t \<cdot> \<delta>) S)"
using term_variants_P.hyps
by simp
have 2: "term_variants_pred P ((map (\<lambda>t. t \<cdot> \<delta>) T) ! i) ((map (\<lambda>t. t \<cdot> \<delta>) S) ! i)"
when "i < length (map (\<lambda>t. t \<cdot> \<delta>) T)" for i
using term_variants_P that
by fastforce
show ?case
using term_variants_pred.term_variants_P[OF 1 2 term_variants_P.hyps(3)]
by fastforce
next
case (term_variants_Fun T S P f)
have 1: "length (map (\<lambda>t. t \<cdot> \<delta>) T) = length (map (\<lambda>t. t \<cdot> \<delta>) S)"
using term_variants_Fun.hyps
by simp
have 2: "term_variants_pred P ((map (\<lambda>t. t \<cdot> \<delta>) T) ! i) ((map (\<lambda>t. t \<cdot> \<delta>) S) ! i)"
when "i < length (map (\<lambda>t. t \<cdot> \<delta>) T)" for i
using term_variants_Fun that
by fastforce
show ?case
using term_variants_pred.term_variants_Fun[OF 1 2]
by fastforce
qed (simp add: term_variants_pred_refl)
lemma term_variants_pred_subst':
fixes t s::"('a,'b) term" and \<delta>::"('a,'b) subst"
assumes "term_variants_pred P (t \<cdot> \<delta>) s"
and "\<forall>x \<in> fv t \<union> fv s. (\<exists>y. \<delta> x = Var y) \<or> (\<exists>f. \<delta> x = Fun f [] \<and> P f = [])"
shows "\<exists>u. term_variants_pred P t u \<and> s = u \<cdot> \<delta>"
using assms
proof (induction P "t \<cdot> \<delta>" s arbitrary: t rule: term_variants_pred.induct)
case (term_variants_Var P x g) thus ?case using term_variants_pred_refl by fast
next
case (term_variants_P T S P g f) show ?case
proof (cases t)
case (Var x) thus ?thesis
using term_variants_P.hyps(4,5) term_variants_P.prems
by fastforce
next
case (Fun h U)
hence 1: "h = f" "T = map (\<lambda>s. s \<cdot> \<delta>) U" "length U = length T"
using term_variants_P.hyps(5) by simp_all
hence 2: "T ! i = U ! i \<cdot> \<delta>" when "i < length T" for i
using that by simp
have "\<forall>x \<in> fv (U ! i) \<union> fv (S ! i). (\<exists>y. \<delta> x = Var y) \<or> (\<exists>f. \<delta> x = Fun f [] \<and> P f = [])"
when "i < length U" for i
using that Fun term_variants_P.prems term_variants_P.hyps(1) 1(3)
by force
hence IH: "\<forall>i < length U. \<exists>u. term_variants_pred P (U ! i) u \<and> S ! i = u \<cdot> \<delta>"
by (metis 1(3) term_variants_P.hyps(3)[OF _ 2])
have "\<exists>V. length U = length V \<and> S = map (\<lambda>v. v \<cdot> \<delta>) V \<and>
(\<forall>i < length U. term_variants_pred P (U ! i) (V ! i))"
using term_variants_P.hyps(1) 1(3) subst_term_list_obtain[OF IH] by metis
then obtain V where V: "length U = length V" "S = map (\<lambda>v. v \<cdot> \<delta>) V"
"\<And>i. i < length U \<Longrightarrow> term_variants_pred P (U ! i) (V ! i)"
by moura
have "term_variants_pred P (Fun f U) (Fun g V)"
by (metis term_variants_pred.term_variants_P[OF V(1,3) term_variants_P.hyps(4)])
moreover have "Fun g S = Fun g V \<cdot> \<delta>" using V(2) by simp
ultimately show ?thesis using term_variants_P.hyps(1,4) Fun 1 by blast
qed
next
case (term_variants_Fun T S P f t) show ?case
proof (cases t)
case (Var x)
hence "T = []" "P f = []" using term_variants_Fun.hyps(4) term_variants_Fun.prems by fastforce+
thus ?thesis using term_variants_pred_refl Var term_variants_Fun.hyps(1,4) by fastforce
next
case (Fun h U)
hence 1: "h = f" "T = map (\<lambda>s. s \<cdot> \<delta>) U" "length U = length T"
using term_variants_Fun.hyps(4) by simp_all
hence 2: "T ! i = U ! i \<cdot> \<delta>" when "i < length T" for i
using that by simp
have "\<forall>x \<in> fv (U ! i) \<union> fv (S ! i). (\<exists>y. \<delta> x = Var y) \<or> (\<exists>f. \<delta> x = Fun f [] \<and> P f = [])"
when "i < length U" for i
using that Fun term_variants_Fun.prems term_variants_Fun.hyps(1) 1(3)
by force
hence IH: "\<forall>i < length U. \<exists>u. term_variants_pred P (U ! i) u \<and> S ! i = u \<cdot> \<delta>"
by (metis 1(3) term_variants_Fun.hyps(3)[OF _ 2 ])
have "\<exists>V. length U = length V \<and> S = map (\<lambda>v. v \<cdot> \<delta>) V \<and>
(\<forall>i < length U. term_variants_pred P (U ! i) (V ! i))"
using term_variants_Fun.hyps(1) 1(3) subst_term_list_obtain[OF IH] by metis
then obtain V where V: "length U = length V" "S = map (\<lambda>v. v \<cdot> \<delta>) V"
"\<And>i. i < length U \<Longrightarrow> term_variants_pred P (U ! i) (V ! i)"
by moura
have "term_variants_pred P (Fun f U) (Fun f V)"
by (metis term_variants_pred.term_variants_Fun[OF V(1,3)])
moreover have "Fun f S = Fun f V \<cdot> \<delta>" using V(2) by simp
ultimately show ?thesis using term_variants_Fun.hyps(1) Fun 1 by blast
qed
qed
lemma term_variants_pred_iff_in_term_variants:
fixes t::"('a,'b) term"
shows "term_variants_pred P t s \<longleftrightarrow> s \<in> set (term_variants P t)"
(is "?A t s \<longleftrightarrow> ?B t s")
proof
define U where "U \<equiv> \<lambda>P (T::('a,'b) term list). product_lists (map (term_variants P) T)"
have a:
"g \<in> set (P f) \<Longrightarrow> set (map (Fun g) (U P T)) \<subseteq> set (term_variants P (Fun f T))"
"set (map (Fun f) (U P T)) \<subseteq> set (term_variants P (Fun f T))"
for f P g and T::"('a,'b) term list"
using term_variants.simps(2)[of P f T]
unfolding U_def Let_def by auto
have b: "\<exists>S \<in> set (U P T). s = Fun f S \<or> (\<exists>g \<in> set (P f). s = Fun g S)"
when "s \<in> set (term_variants P (Fun f T))" for P T f s
using that by (cases "P f") (auto simp add: U_def Let_def)
have c: "length T = length S" when "S \<in> set (U P T)" for S P T
using that unfolding U_def
by (simp add: in_set_product_lists_length)
show "?A t s \<Longrightarrow> ?B t s"
proof (induction P t s rule: term_variants_pred.induct)
case (term_variants_P T S P g f)
note hyps = term_variants_P.hyps
note IH = term_variants_P.IH
have "S \<in> set (U P T)"
using IH hyps(1) product_lists_in_set_nth'[of _ S]
unfolding U_def by simp
thus ?case using a(1)[of _ P, OF hyps(3)] by auto
next
case (term_variants_Fun T S P f)
note hyps = term_variants_Fun.hyps
note IH = term_variants_Fun.IH
have "S \<in> set (U P T)"
using IH hyps(1) product_lists_in_set_nth'[of _ S]
unfolding U_def by simp
thus ?case using a(2)[of f P T] by (cases "P f") auto
qed (simp add: term_variants_Var)
show "?B t s \<Longrightarrow> ?A t s"
proof (induction P t arbitrary: s rule: term_variants.induct)
case (2 P f T)
obtain S where S:
"s = Fun f S \<or> (\<exists>g \<in> set (P f). s = Fun g S)"
"S \<in> set (U P T)" "length T = length S"
using c b[OF "2.prems"] by moura
have "\<forall>i < length T. term_variants_pred P (T ! i) (S ! i)"
using "2.IH" S product_lists_in_set_nth by (fastforce simp add: U_def)
thus ?case using S by (auto intro: term_variants_pred.intros)
qed (simp add: term_variants_Var)
qed
lemma term_variants_pred_finite:
"finite {s. term_variants_pred P t s}"
using term_variants_pred_iff_in_term_variants[of P t]
by simp
lemma term_variants_pred_fv_eq:
assumes "term_variants_pred P s t"
shows "fv s = fv t"
using assms
by (induct rule: term_variants_pred.induct)
(metis, metis fv_eq_FunI, metis fv_eq_FunI)
lemma (in intruder_model) term_variants_pred_wf_trms:
assumes "term_variants_pred P s t"
and "\<And>f g. g \<in> set (P f) \<Longrightarrow> arity f = arity g"
and "wf\<^sub>t\<^sub>r\<^sub>m s"
shows "wf\<^sub>t\<^sub>r\<^sub>m t"
using assms
apply (induction rule: term_variants_pred.induct, simp)
by (metis (no_types) wf_trmI wf_trm_arity in_set_conv_nth wf_trm_param_idx)+
lemma term_variants_pred_funs_term:
assumes "term_variants_pred P s t"
and "f \<in> funs_term t"
shows "f \<in> funs_term s \<or> (\<exists>g \<in> funs_term s. f \<in> set (P g))"
using assms
proof (induction rule: term_variants_pred.induct)
case (term_variants_P T S P g h) thus ?case
proof (cases "f = g")
case False
then obtain s where "s \<in> set S" "f \<in> funs_term s"
using funs_term_subterms_eq(1)[of "Fun g S"] term_variants_P.prems by auto
thus ?thesis
using term_variants_P.IH term_variants_P.hyps(1) in_set_conv_nth[of s S] by force
qed simp
next
case (term_variants_Fun T S P h) thus ?case
proof (cases "f = h")
case False
then obtain s where "s \<in> set S" "f \<in> funs_term s"
using funs_term_subterms_eq(1)[of "Fun h S"] term_variants_Fun.prems by auto
thus ?thesis
using term_variants_Fun.IH term_variants_Fun.hyps(1) in_set_conv_nth[of s S] by force
qed simp
qed fast
end

View File

@ -0,0 +1,966 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Transactions.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>Protocol Transactions\<close>
theory Transactions
imports
Stateful_Protocol_Composition_and_Typing.Typed_Model
Stateful_Protocol_Composition_and_Typing.Labeled_Stateful_Strands
begin
subsection \<open>Definitions\<close>
datatype 'b prot_atom =
is_Atom: Atom 'b
| Value
| SetType
| AttackType
| Bottom
| OccursSecType
datatype ('a,'b,'c) prot_fun =
Fu (the_Fu: 'a)
| Set (the_Set: 'c)
| Val (the_Val: "nat \<times> bool")
| Abs (the_Abs: "'c set")
| Pair
| Attack nat
| PubConstAtom 'b nat
| PubConstSetType nat
| PubConstAttackType nat
| PubConstBottom nat
| PubConstOccursSecType nat
| OccursFact
| OccursSec
definition "is_Fun_Set t \<equiv> is_Fun t \<and> args t = [] \<and> is_Set (the_Fun t)"
abbreviation occurs where
"occurs t \<equiv> Fun OccursFact [Fun OccursSec [], t]"
type_synonym ('a,'b,'c) prot_term_type = "(('a,'b,'c) prot_fun,'b prot_atom) term_type"
type_synonym ('a,'b,'c) prot_var = "('a,'b,'c) prot_term_type \<times> nat"
type_synonym ('a,'b,'c) prot_term = "(('a,'b,'c) prot_fun,('a,'b,'c) prot_var) term"
type_synonym ('a,'b,'c) prot_terms = "('a,'b,'c) prot_term set"
type_synonym ('a,'b,'c) prot_subst = "(('a,'b,'c) prot_fun, ('a,'b,'c) prot_var) subst"
type_synonym ('a,'b,'c,'d) prot_strand_step =
"(('a,'b,'c) prot_fun, ('a,'b,'c) prot_var, 'd) labeled_stateful_strand_step"
type_synonym ('a,'b,'c,'d) prot_strand = "('a,'b,'c,'d) prot_strand_step list"
type_synonym ('a,'b,'c,'d) prot_constr = "('a,'b,'c,'d) prot_strand_step list"
datatype ('a,'b,'c,'d) prot_transaction =
Transaction
(transaction_fresh: "('a,'b,'c) prot_var list")
(transaction_receive: "('a,'b,'c,'d) prot_strand")
(transaction_selects: "('a,'b,'c,'d) prot_strand")
(transaction_checks: "('a,'b,'c,'d) prot_strand")
(transaction_updates: "('a,'b,'c,'d) prot_strand")
(transaction_send: "('a,'b,'c,'d) prot_strand")
definition transaction_strand where
"transaction_strand T \<equiv>
transaction_receive T@transaction_selects T@transaction_checks T@
transaction_updates T@transaction_send T"
fun transaction_proj where
"transaction_proj l (Transaction A B C D E F) = (
let f = proj l
in Transaction A (f B) (f C) (f D) (f E) (f F))"
fun transaction_star_proj where
"transaction_star_proj (Transaction A B C D E F) = (
let f = filter is_LabelS
in Transaction A (f B) (f C) (f D) (f E) (f F))"
abbreviation fv_transaction where
"fv_transaction T \<equiv> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)"
abbreviation bvars_transaction where
"bvars_transaction T \<equiv> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)"
abbreviation vars_transaction where
"vars_transaction T \<equiv> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)"
abbreviation trms_transaction where
"trms_transaction T \<equiv> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)"
abbreviation setops_transaction where
"setops_transaction T \<equiv> setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))"
definition wellformed_transaction where
"wellformed_transaction T \<equiv>
list_all is_Receive (unlabel (transaction_receive T)) \<and>
list_all is_Assignment (unlabel (transaction_selects T)) \<and>
list_all is_Check (unlabel (transaction_checks T)) \<and>
list_all is_Update (unlabel (transaction_updates T)) \<and>
list_all is_Send (unlabel (transaction_send T)) \<and>
set (transaction_fresh T) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<and>
set (transaction_fresh T) \<inter> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {} \<and>
set (transaction_fresh T) \<inter> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = {} \<and>
fv_transaction T \<inter> bvars_transaction T = {} \<and>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<and>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T)
\<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<and>
(\<forall>x \<in> set (unlabel (transaction_selects T)).
is_Equality x \<longrightarrow> fv (the_rhs x) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))"
type_synonym ('a,'b,'c,'d) prot = "('a,'b,'c,'d) prot_transaction list"
abbreviation Var_Value_term ("\<langle>_\<rangle>\<^sub>v") where
"\<langle>n\<rangle>\<^sub>v \<equiv> Var (Var Value, n)::('a,'b,'c) prot_term"
abbreviation Fun_Fu_term ("\<langle>_ _\<rangle>\<^sub>t") where
"\<langle>f T\<rangle>\<^sub>t \<equiv> Fun (Fu f) T::('a,'b,'c) prot_term"
abbreviation Fun_Fu_const_term ("\<langle>_\<rangle>\<^sub>c") where
"\<langle>c\<rangle>\<^sub>c \<equiv> Fun (Fu c) []::('a,'b,'c) prot_term"
abbreviation Fun_Set_const_term ("\<langle>_\<rangle>\<^sub>s") where
"\<langle>f\<rangle>\<^sub>s \<equiv> Fun (Set f) []::('a,'b,'c) prot_term"
abbreviation Fun_Abs_const_term ("\<langle>_\<rangle>\<^sub>a") where
"\<langle>a\<rangle>\<^sub>a \<equiv> Fun (Abs a) []::('a,'b,'c) prot_term"
abbreviation Fun_Attack_const_term ("attack\<langle>_\<rangle>") where
"attack\<langle>n\<rangle> \<equiv> Fun (Attack n) []::('a,'b,'c) prot_term"
abbreviation prot_transaction1 ("transaction\<^sub>1 _ _ new _ _ _") where
"transaction\<^sub>1 (S1::('a,'b,'c,'d) prot_strand) S2 new (B::('a,'b,'c) prot_term list) S3 S4
\<equiv> Transaction (map the_Var B) S1 [] S2 S3 S4"
abbreviation prot_transaction2 ("transaction\<^sub>2 _ _ _ _") where
"transaction\<^sub>2 (S1::('a,'b,'c,'d) prot_strand) S2 S3 S4
\<equiv> Transaction [] S1 [] S2 S3 S4"
subsection \<open>Lemmata\<close>
lemma prot_atom_UNIV:
"(UNIV::'b prot_atom set) = range Atom \<union> {Value, SetType, AttackType, Bottom, OccursSecType}"
proof -
have "a \<in> range Atom \<or> a = Value \<or> a = SetType \<or> a = AttackType \<or> a = Bottom \<or> a = OccursSecType"
for a::"'b prot_atom"
by (cases a) auto
thus ?thesis by auto
qed
instance prot_atom::(finite) finite
by intro_classes (simp add: prot_atom_UNIV)
instantiation prot_atom::(enum) enum
begin
definition "enum_prot_atom == map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType]"
definition "enum_all_prot_atom P == list_all P (map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType])"
definition "enum_ex_prot_atom P == list_ex P (map Atom enum_class.enum@[Value, SetType, AttackType, Bottom, OccursSecType])"
instance
proof intro_classes
have *: "set (map Atom (enum_class.enum::'a list)) = range Atom"
"distinct (enum_class.enum::'a list)"
using UNIV_enum enum_distinct by auto
show "(UNIV::'a prot_atom set) = set enum_class.enum"
using *(1) by (simp add: prot_atom_UNIV enum_prot_atom_def)
have "set (map Atom enum_class.enum) \<inter> set [Value, SetType, AttackType, Bottom, OccursSecType] = {}" by auto
moreover have "inj_on Atom (set (enum_class.enum::'a list))" unfolding inj_on_def by auto
hence "distinct (map Atom (enum_class.enum::'a list))" by (metis *(2) distinct_map)
ultimately show "distinct (enum_class.enum::'a prot_atom list)" by (simp add: enum_prot_atom_def)
have "Ball UNIV P \<longleftrightarrow> Ball (range Atom) P \<and> Ball {Value, SetType, AttackType, Bottom, OccursSecType} P"
for P::"'a prot_atom \<Rightarrow> bool"
by (metis prot_atom_UNIV UNIV_I UnE)
thus "enum_class.enum_all P = Ball (UNIV::'a prot_atom set) P" for P
using *(1) Ball_set[of "map Atom enum_class.enum" P]
by (auto simp add: enum_all_prot_atom_def)
have "Bex UNIV P \<longleftrightarrow> Bex (range Atom) P \<or> Bex {Value, SetType, AttackType, Bottom, OccursSecType} P"
for P::"'a prot_atom \<Rightarrow> bool"
by (metis prot_atom_UNIV UNIV_I UnE)
thus "enum_class.enum_ex P = Bex (UNIV::'a prot_atom set) P" for P
using *(1) Bex_set[of "map Atom enum_class.enum" P]
by (auto simp add: enum_ex_prot_atom_def)
qed
end
lemma wellformed_transaction_cases:
assumes "wellformed_transaction T"
shows
"(l,x) \<in> set (transaction_receive T) \<Longrightarrow> \<exists>t. x = receive\<langle>t\<rangle>" (is "?A \<Longrightarrow> ?A'")
"(l,x) \<in> set (transaction_selects T) \<Longrightarrow>
(\<exists>t s. x = \<langle>t := s\<rangle>) \<or> (\<exists>t s. x = select\<langle>t,s\<rangle>)" (is "?B \<Longrightarrow> ?B'")
"(l,x) \<in> set (transaction_checks T) \<Longrightarrow>
(\<exists>t s. x = \<langle>t == s\<rangle>) \<or> (\<exists>t s. x = \<langle>t in s\<rangle>) \<or> (\<exists>X F G. x = \<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>)" (is "?C \<Longrightarrow> ?C'")
"(l,x) \<in> set (transaction_updates T) \<Longrightarrow>
(\<exists>t s. x = insert\<langle>t,s\<rangle>) \<or> (\<exists>t s. x = delete\<langle>t,s\<rangle>)" (is "?D \<Longrightarrow> ?D'")
"(l,x) \<in> set (transaction_send T) \<Longrightarrow> \<exists>t. x = send\<langle>t\<rangle>" (is "?E \<Longrightarrow> ?E'")
proof -
have a:
"list_all is_Receive (unlabel (transaction_receive T))"
"list_all is_Assignment (unlabel (transaction_selects T))"
"list_all is_Check (unlabel (transaction_checks T))"
"list_all is_Update (unlabel (transaction_updates T))"
"list_all is_Send (unlabel (transaction_send T))"
using assms unfolding wellformed_transaction_def by metis+
note b = Ball_set unlabel_in
note c = stateful_strand_step.collapse
show "?A \<Longrightarrow> ?A'" by (metis (mono_tags, lifting) a(1) b c(2))
show "?B \<Longrightarrow> ?B'" by (metis (mono_tags, lifting) a(2) b c(3,6))
show "?C \<Longrightarrow> ?C'" by (metis (mono_tags, lifting) a(3) b c(3,6,7))
show "?D \<Longrightarrow> ?D'" by (metis (mono_tags, lifting) a(4) b c(4,5))
show "?E \<Longrightarrow> ?E'" by (metis (mono_tags, lifting) a(5) b c(1))
qed
lemma wellformed_transaction_unlabel_cases:
assumes "wellformed_transaction T"
shows
"x \<in> set (unlabel (transaction_receive T)) \<Longrightarrow> \<exists>t. x = receive\<langle>t\<rangle>" (is "?A \<Longrightarrow> ?A'")
"x \<in> set (unlabel (transaction_selects T)) \<Longrightarrow>
(\<exists>t s. x = \<langle>t := s\<rangle>) \<or> (\<exists>t s. x = select\<langle>t,s\<rangle>)" (is "?B \<Longrightarrow> ?B'")
"x \<in> set (unlabel (transaction_checks T)) \<Longrightarrow>
(\<exists>t s. x = \<langle>t == s\<rangle>) \<or> (\<exists>t s. x = \<langle>t in s\<rangle>) \<or> (\<exists>X F G. x = \<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle>)"
(is "?C \<Longrightarrow> ?C'")
"x \<in> set (unlabel (transaction_updates T)) \<Longrightarrow>
(\<exists>t s. x = insert\<langle>t,s\<rangle>) \<or> (\<exists>t s. x = delete\<langle>t,s\<rangle>)" (is "?D \<Longrightarrow> ?D'")
"x \<in> set (unlabel (transaction_send T)) \<Longrightarrow> \<exists>t. x = send\<langle>t\<rangle>" (is "?E \<Longrightarrow> ?E'")
proof -
have a:
"list_all is_Receive (unlabel (transaction_receive T))"
"list_all is_Assignment (unlabel (transaction_selects T))"
"list_all is_Check (unlabel (transaction_checks T))"
"list_all is_Update (unlabel (transaction_updates T))"
"list_all is_Send (unlabel (transaction_send T))"
using assms unfolding wellformed_transaction_def by metis+
note b = Ball_set
note c = stateful_strand_step.collapse
show "?A \<Longrightarrow> ?A'" by (metis (mono_tags, lifting) a(1) b c(2))
show "?B \<Longrightarrow> ?B'" by (metis (mono_tags, lifting) a(2) b c(3,6))
show "?C \<Longrightarrow> ?C'" by (metis (mono_tags, lifting) a(3) b c(3,6,7))
show "?D \<Longrightarrow> ?D'" by (metis (mono_tags, lifting) a(4) b c(4,5))
show "?E \<Longrightarrow> ?E'" by (metis (mono_tags, lifting) a(5) b c(1))
qed
lemma transaction_strand_subsets[simp]:
"set (transaction_receive T) \<subseteq> set (transaction_strand T)"
"set (transaction_selects T) \<subseteq> set (transaction_strand T)"
"set (transaction_checks T) \<subseteq> set (transaction_strand T)"
"set (transaction_updates T) \<subseteq> set (transaction_strand T)"
"set (transaction_send T) \<subseteq> set (transaction_strand T)"
"set (unlabel (transaction_receive T)) \<subseteq> set (unlabel (transaction_strand T))"
"set (unlabel (transaction_selects T)) \<subseteq> set (unlabel (transaction_strand T))"
"set (unlabel (transaction_checks T)) \<subseteq> set (unlabel (transaction_strand T))"
"set (unlabel (transaction_updates T)) \<subseteq> set (unlabel (transaction_strand T))"
"set (unlabel (transaction_send T)) \<subseteq> set (unlabel (transaction_strand T))"
unfolding transaction_strand_def unlabel_def by force+
lemma transaction_strand_subst_subsets[simp]:
"set (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<subseteq> set (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
"set (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<subseteq> set (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
"set (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<subseteq> set (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
"set (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<subseteq> set (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
"set (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<subseteq> set (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
"set (unlabel (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) \<subseteq> set (unlabel (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
"set (unlabel (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) \<subseteq> set (unlabel (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
"set (unlabel (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) \<subseteq> set (unlabel (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
"set (unlabel (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) \<subseteq> set (unlabel (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
"set (unlabel (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) \<subseteq> set (unlabel (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
unfolding transaction_strand_def unlabel_def subst_apply_labeled_stateful_strand_def by force+
lemma transaction_dual_subst_unfold:
"unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)) =
unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))@
unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))@
unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))@
unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))@
unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
by (simp add: transaction_strand_def unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append)
lemma trms_transaction_unfold:
"trms_transaction T =
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<union>
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<union> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union>
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
by (metis trms\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def)
lemma trms_transaction_subst_unfold:
"trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) =
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
by (metis trms\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append)
lemma vars_transaction_unfold:
"vars_transaction T =
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<union>
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<union> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union>
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
by (metis vars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def)
lemma vars_transaction_subst_unfold:
"vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) =
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
by (metis vars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append)
lemma fv_transaction_unfold:
"fv_transaction T =
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<union>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
by (metis fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def)
lemma fv_transaction_subst_unfold:
"fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) =
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
by (metis fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append)
lemma fv_wellformed_transaction_unfold:
assumes "wellformed_transaction T"
shows "fv_transaction T =
fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<union> set (transaction_fresh T)"
proof -
let ?A = "set (transaction_fresh T)"
let ?B = "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T)"
let ?C = "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
let ?D = "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)"
let ?E = "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)"
let ?F = "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)"
have "?A \<subseteq> ?B \<union> ?C" "?A \<inter> ?D = {}" "?A \<inter> ?E = {}" "?F \<subseteq> ?D \<union> ?E" "?B \<union> ?C - ?A \<subseteq> ?D \<union> ?E"
using assms unfolding wellformed_transaction_def by fast+
thus ?thesis using fv_transaction_unfold by blast
qed
lemma bvars_transaction_unfold:
"bvars_transaction T =
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \<union>
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<union> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union>
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
by (metis bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def)
lemma bvars_transaction_subst_unfold:
"bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) =
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) \<union>
bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
by (metis bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append append_assoc transaction_strand_def subst_lsst_append)
lemma bvars_wellformed_transaction_unfold:
assumes "wellformed_transaction T"
shows "bvars_transaction T = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?A)
and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" (is ?B)
and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = {}" (is ?C)
and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) = {}" (is ?D)
and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) = {}" (is ?E)
proof -
have 0: "list_all is_Receive (unlabel (transaction_receive T))"
"list_all is_Assignment (unlabel (transaction_selects T))"
"list_all is_Update (unlabel (transaction_updates T))"
"list_all is_Send (unlabel (transaction_send T))"
using assms unfolding wellformed_transaction_def by metis+
have "filter is_NegChecks (unlabel (transaction_receive T)) = []"
"filter is_NegChecks (unlabel (transaction_selects T)) = []"
"filter is_NegChecks (unlabel (transaction_updates T)) = []"
"filter is_NegChecks (unlabel (transaction_send T)) = []"
using list_all_filter_nil[OF 0(1), of is_NegChecks]
list_all_filter_nil[OF 0(2), of is_NegChecks]
list_all_filter_nil[OF 0(3), of is_NegChecks]
list_all_filter_nil[OF 0(4), of is_NegChecks]
stateful_strand_step.distinct_disc(11,21,29,35,39,41)
by blast+
thus ?A ?B ?C ?D ?E
using bvars_transaction_unfold[of T]
bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_receive T)"]
bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_selects T)"]
bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_updates T)"]
bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks[of "unlabel (transaction_send T)"]
by (metis bvars\<^sub>s\<^sub>s\<^sub>t_def UnionE emptyE list.set(1) list.simps(8) subsetI subset_Un_eq sup_commute)+
qed
lemma transaction_strand_memberD[dest]:
assumes "x \<in> set (transaction_strand T)"
shows "x \<in> set (transaction_receive T) \<or> x \<in> set (transaction_selects T) \<or>
x \<in> set (transaction_checks T) \<or> x \<in> set (transaction_updates T) \<or>
x \<in> set (transaction_send T)"
using assms by (simp add: transaction_strand_def)
lemma transaction_strand_unlabel_memberD[dest]:
assumes "x \<in> set (unlabel (transaction_strand T))"
shows "x \<in> set (unlabel (transaction_receive T)) \<or> x \<in> set (unlabel (transaction_selects T)) \<or>
x \<in> set (unlabel (transaction_checks T)) \<or> x \<in> set (unlabel (transaction_updates T)) \<or>
x \<in> set (unlabel (transaction_send T))"
using assms by (simp add: unlabel_def transaction_strand_def)
lemma wellformed_transaction_strand_memberD[dest]:
assumes "wellformed_transaction T" and "(l,x) \<in> set (transaction_strand T)"
shows
"x = receive\<langle>t\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_receive T)" (is "?A \<Longrightarrow> ?A'")
"x = select\<langle>t,s\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_selects T)" (is "?B \<Longrightarrow> ?B'")
"x = \<langle>t == s\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_checks T)" (is "?C \<Longrightarrow> ?C'")
"x = \<langle>t in s\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_checks T)" (is "?D \<Longrightarrow> ?D'")
"x = \<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_checks T)" (is "?E \<Longrightarrow> ?E'")
"x = insert\<langle>t,s\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_updates T)" (is "?F \<Longrightarrow> ?F'")
"x = delete\<langle>t,s\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_updates T)" (is "?G \<Longrightarrow> ?G'")
"x = send\<langle>t\<rangle> \<Longrightarrow> (l,x) \<in> set (transaction_send T)" (is "?H \<Longrightarrow> ?H'")
proof -
have "(l,x) \<in> set (transaction_receive T) \<or> (l,x) \<in> set (transaction_selects T) \<or>
(l,x) \<in> set (transaction_checks T) \<or> (l,x) \<in> set (transaction_updates T) \<or>
(l,x) \<in> set (transaction_send T)"
using assms(2) by auto
thus "?A \<Longrightarrow> ?A'" "?B \<Longrightarrow> ?B'" "?C \<Longrightarrow> ?C'" "?D \<Longrightarrow> ?D'"
"?E \<Longrightarrow> ?E'" "?F \<Longrightarrow> ?F'" "?G \<Longrightarrow> ?G'" "?H \<Longrightarrow> ?H'"
using wellformed_transaction_cases[OF assms(1)] by fast+
qed
lemma wellformed_transaction_strand_unlabel_memberD[dest]:
assumes "wellformed_transaction T" and "x \<in> set (unlabel (transaction_strand T))"
shows
"x = receive\<langle>t\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_receive T))" (is "?A \<Longrightarrow> ?A'")
"x = select\<langle>t,s\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_selects T))" (is "?B \<Longrightarrow> ?B'")
"x = \<langle>t == s\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_checks T))" (is "?C \<Longrightarrow> ?C'")
"x = \<langle>t in s\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_checks T))" (is "?D \<Longrightarrow> ?D'")
"x = \<forall>X\<langle>\<or>\<noteq>: F \<or>\<notin>: G\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_checks T))" (is "?E \<Longrightarrow> ?E'")
"x = insert\<langle>t,s\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_updates T))" (is "?F \<Longrightarrow> ?F'")
"x = delete\<langle>t,s\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_updates T))" (is "?G \<Longrightarrow> ?G'")
"x = send\<langle>t\<rangle> \<Longrightarrow> x \<in> set (unlabel (transaction_send T))" (is "?H \<Longrightarrow> ?H'")
proof -
have "x \<in> set (unlabel (transaction_receive T)) \<or> x \<in> set (unlabel (transaction_selects T)) \<or>
x \<in> set (unlabel (transaction_checks T)) \<or> x \<in> set (unlabel (transaction_updates T)) \<or>
x \<in> set (unlabel (transaction_send T))"
using assms(2) by auto
thus "?A \<Longrightarrow> ?A'" "?B \<Longrightarrow> ?B'" "?C \<Longrightarrow> ?C'" "?D \<Longrightarrow> ?D'"
"?E \<Longrightarrow> ?E'" "?F \<Longrightarrow> ?F'" "?G \<Longrightarrow> ?G'" "?H \<Longrightarrow> ?H'"
using wellformed_transaction_unlabel_cases[OF assms(1)] by fast+
qed
lemma wellformed_transaction_send_receive_trm_cases:
assumes T: "wellformed_transaction T"
shows "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<Longrightarrow> receive\<langle>t\<rangle> \<in> set (unlabel (transaction_receive T))"
and "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<Longrightarrow> send\<langle>t\<rangle> \<in> set (unlabel (transaction_send T))"
using wellformed_transaction_unlabel_cases(1,5)[OF T]
trms\<^sub>s\<^sub>s\<^sub>t_in[of t "unlabel (transaction_receive T)"]
trms\<^sub>s\<^sub>s\<^sub>t_in[of t "unlabel (transaction_send T)"]
by fastforce+
lemma wellformed_transaction_send_receive_subst_trm_cases:
assumes T: "wellformed_transaction T"
shows "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta> \<Longrightarrow> receive\<langle>t\<rangle> \<in> set (unlabel (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
and "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta> \<Longrightarrow> send\<langle>t\<rangle> \<in> set (unlabel (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
proof -
assume "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta>"
then obtain s where s: "s \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" "t = s \<cdot> \<theta>"
by blast
hence "receive\<langle>s\<rangle> \<in> set (unlabel (transaction_receive T))"
using wellformed_transaction_send_receive_trm_cases(1)[OF T] by simp
thus "receive\<langle>t\<rangle> \<in> set (unlabel (transaction_receive T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
by (metis s(2) unlabel_subst[of _ \<theta>] stateful_strand_step_subst_inI(2))
next
assume "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<theta>"
then obtain s where s: "s \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t = s \<cdot> \<theta>"
by blast
hence "send\<langle>s\<rangle> \<in> set (unlabel (transaction_send T))"
using wellformed_transaction_send_receive_trm_cases(2)[OF T] by simp
thus "send\<langle>t\<rangle> \<in> set (unlabel (transaction_send T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>))"
by (metis s(2) unlabel_subst[of _ \<theta>] stateful_strand_step_subst_inI(1))
qed
lemma wellformed_transaction_send_receive_fv_subset:
assumes T: "wellformed_transaction T"
shows "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<Longrightarrow> fv t \<subseteq> fv_transaction T" (is "?A \<Longrightarrow> ?A'")
and "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<Longrightarrow> fv t \<subseteq> fv_transaction T" (is "?B \<Longrightarrow> ?B'")
proof -
have "t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<Longrightarrow> receive\<langle>t\<rangle> \<in> set (unlabel (transaction_strand T))"
"t \<in> trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \<Longrightarrow> send\<langle>t\<rangle> \<in> set (unlabel (transaction_strand T))"
using wellformed_transaction_send_receive_trm_cases[OF T, of t]
unfolding transaction_strand_def by force+
thus "?A \<Longrightarrow> ?A'" "?B \<Longrightarrow> ?B'" by (induct "transaction_strand T") auto
qed
lemma dual_wellformed_transaction_ident_cases[dest]:
"list_all is_Assignment (unlabel S) \<Longrightarrow> dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S"
"list_all is_Check (unlabel S) \<Longrightarrow> dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S"
"list_all is_Update (unlabel S) \<Longrightarrow> dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S = S"
proof (induction S)
case (Cons s S)
obtain l x where s: "s = (l,x)" by moura
{ case 1 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto }
{ case 2 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto }
{ case 3 thus ?case using Cons s unfolding unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (cases x) auto }
qed simp_all
lemma wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t:
fixes T::"('a, 'b, 'c, 'd) prot_transaction"
assumes T: "wellformed_transaction T"
shows "wf'\<^sub>s\<^sub>s\<^sub>t (set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" (is ?A)
and "fv_transaction T \<inter> bvars_transaction T = {}" (is ?B)
and "set (transaction_fresh T) \<inter> bvars_transaction T = {}" (is ?C)
proof -
define T1 where "T1 \<equiv> unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))"
define T2 where "T2 \<equiv> unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T))"
define T3 where "T3 \<equiv> unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T))"
define T4 where "T4 \<equiv> unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T))"
define T5 where "T5 \<equiv> unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))"
define X where "X \<equiv> set (transaction_fresh T)"
define Y where "Y \<equiv> X \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1"
define Z where "Z \<equiv> Y \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T2"
define f where "f \<equiv> \<lambda>S::(('a,'b,'c) prot_fun, ('a,'b,'c) prot_var) stateful_strand.
\<Union>((\<lambda>x. case x of
Receive t \<Rightarrow> fv t
| Equality Assign _ t' \<Rightarrow> fv t'
| Insert t t' \<Rightarrow> fv t \<union> fv t'
| _ \<Rightarrow> {}) ` set S)"
note defs1 = T1_def T2_def T3_def T4_def T5_def
note defs2 = X_def Y_def Z_def
note defs3 = f_def
have 0: "wf'\<^sub>s\<^sub>s\<^sub>t V (S @ S')"
when "wf'\<^sub>s\<^sub>s\<^sub>t V S" "f S' \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \<union> V" for V S S'
by (metis that wf\<^sub>s\<^sub>s\<^sub>t_append_suffix' f_def)
have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = T1@T2@T3@T4@T5"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unlabel_append unfolding transaction_strand_def defs1 by simp
have 2:
"\<forall>x \<in> set T1. is_Send x" "\<forall>x \<in> set T2. is_Assignment x" "\<forall>x \<in> set T3. is_Check x"
"\<forall>x \<in> set T4. is_Update x" "\<forall>x \<in> set T5. is_Receive x"
"fv\<^sub>s\<^sub>s\<^sub>t T3 \<subseteq> fv\<^sub>s\<^sub>s\<^sub>t T1 \<union> fv\<^sub>s\<^sub>s\<^sub>t T2" "fv\<^sub>s\<^sub>s\<^sub>t T4 \<union> fv\<^sub>s\<^sub>s\<^sub>t T5 \<subseteq> X \<union> fv\<^sub>s\<^sub>s\<^sub>t T1 \<union> fv\<^sub>s\<^sub>s\<^sub>t T2"
"X \<inter> fv\<^sub>s\<^sub>s\<^sub>t T1 = {}" "X \<inter> fv\<^sub>s\<^sub>s\<^sub>t T2 = {}"
"\<forall>x \<in> set T2. is_Equality x \<longrightarrow> fv (the_rhs x) \<subseteq> fv\<^sub>s\<^sub>s\<^sub>t T1"
using T unfolding defs1 defs2 wellformed_transaction_def
by (auto simp add: Ball_set dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq simp del: fv\<^sub>s\<^sub>s\<^sub>t_def)
have 3: "wf'\<^sub>s\<^sub>s\<^sub>t X T1" using 2(1)
proof (induction T1 arbitrary: X)
case (Cons s T)
obtain t where "s = send\<langle>t\<rangle>" using Cons.prems by (cases s) moura+
thus ?case using Cons by auto
qed simp
have 4: "f T1 = {}" "fv\<^sub>s\<^sub>s\<^sub>t T1 = wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1" using 2(1)
proof (induction T1)
case (Cons s T)
{ case 1 thus ?case using Cons unfolding defs3 by (cases s) auto }
{ case 2 thus ?case using Cons unfolding defs3 wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def by (cases s) auto }
qed (simp_all add: defs3 wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def)
have 5: "f T2 \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1" "fv\<^sub>s\<^sub>s\<^sub>t T2 = f T2 \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T2" using 2(2,10)
proof (induction T2)
case (Cons s T)
{ case 1 thus ?case using Cons
proof (cases s)
case (Equality ac t t') thus ?thesis using 1 Cons 4(2) unfolding defs3 by (cases ac) auto
qed (simp_all add: defs3)
}
{ case 2 thus ?case using Cons
proof (cases s)
case (Equality ac t t')
hence "ac = Assign" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = fv t' \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" "f (s#T) = fv t' \<union> f T"
using 2 unfolding defs3 by auto
moreover have "fv\<^sub>s\<^sub>s\<^sub>t T = f T \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T" using Cons.IH(2) 2 by auto
ultimately show ?thesis unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def by auto
next
case (InSet ac t t')
hence "ac = Assign" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" "f (s#T) = f T"
using 2 unfolding defs3 by auto
moreover have "fv\<^sub>s\<^sub>s\<^sub>t T = f T \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T" using Cons.IH(2) 2 by auto
ultimately show ?thesis unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def by auto
qed (simp_all add: defs3)
}
qed (simp_all add: defs3 wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def)
have "f T \<subseteq> fv\<^sub>s\<^sub>s\<^sub>t T" for T
proof
fix x show "x \<in> f T \<Longrightarrow> x \<in> fv\<^sub>s\<^sub>s\<^sub>t T"
proof (induction T)
case (Cons s T) thus ?case
proof (cases "x \<in> f T")
case False thus ?thesis
using Cons.prems unfolding defs3 fv\<^sub>s\<^sub>s\<^sub>t_def
by (auto split: stateful_strand_step.splits poscheckvariant.splits)
qed auto
qed (simp add: defs3 fv\<^sub>s\<^sub>s\<^sub>t_def)
qed
hence 6:
"f T3 \<subseteq> X \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1 \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T2"
"f T4 \<subseteq> X \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1 \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T2"
"f T5 \<subseteq> X \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1 \<union> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T2"
using 2(6,7) 4 5 by blast+
have 7:
"wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T3 = {}"
"wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T4 = {}"
"wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T5 = {}"
using 2(3,4,5) unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def
by (auto split: stateful_strand_step.splits)
have 8:
"f T2 \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t T1 \<union> X"
"f T3 \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (T1@T2) \<union> X"
"f T4 \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t ((T1@T2)@T3) \<union> X"
"f T5 \<subseteq> wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (((T1@T2)@T3)@T4) \<union> X"
using 4(1) 5(1) 6 7 wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_append[of T1 T2]
wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_append[of "T1@T2" T3]
wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_append[of "(T1@T2)@T3" T4]
by blast+
have "wf'\<^sub>s\<^sub>s\<^sub>t X (T1@T2@T3@T4@T5)"
using 0[OF 0[OF 0[OF 0[OF 3 8(1)] 8(2)] 8(3)] 8(4)]
unfolding Y_def Z_def by simp
thus ?A using 1 unfolding defs1 defs2 by simp
have "set (transaction_fresh T) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
"fv_transaction T \<inter> bvars_transaction T = {}"
using T unfolding wellformed_transaction_def by fast+
thus ?B ?C using fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast+
qed
lemma dual_wellformed_transaction_ident_cases'[dest]:
assumes "wellformed_transaction T"
shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = transaction_selects T"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = transaction_checks T"
"dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) = transaction_updates T"
using assms unfolding wellformed_transaction_def by auto
lemma dual_transaction_strand:
assumes "wellformed_transaction T"
shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) =
dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)@transaction_selects T@transaction_checks T@
transaction_updates T@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
using dual_wellformed_transaction_ident_cases'[OF assms] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append
unfolding transaction_strand_def by metis
lemma dual_unlabel_transaction_strand:
assumes "wellformed_transaction T"
shows "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) =
(unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)))@(unlabel (transaction_selects T))@
(unlabel (transaction_checks T))@(unlabel (transaction_updates T))@
(unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)))"
using dual_transaction_strand[OF assms] by (simp add: unlabel_def)
lemma dual_transaction_strand_subst:
assumes "wellformed_transaction T"
shows "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) =
(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)@transaction_selects T@transaction_checks T@
transaction_updates T@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>"
proof -
have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by metis
thus ?thesis using dual_transaction_strand[OF assms] by argo
qed
lemma dual_transaction_ik_is_transaction_send:
assumes "wellformed_transaction T"
shows "ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))) = trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T))"
(is "?A = ?B")
proof -
{ fix t assume "t \<in> ?A"
hence "receive\<langle>t\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def)
hence "send\<langle>t\<rangle> \<in> set (unlabel (transaction_strand T))"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by metis
hence "t \<in> ?B" using wellformed_transaction_strand_unlabel_memberD(8)[OF assms] by force
} moreover {
fix t assume "t \<in> ?B"
hence "send\<langle>t\<rangle> \<in> set (unlabel (transaction_send T))"
using wellformed_transaction_unlabel_cases(5)[OF assms] by fastforce
hence "receive\<langle>t\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)))"
using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by metis
hence "receive\<langle>t\<rangle> \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))"
using dual_unlabel_transaction_strand[OF assms] by simp
hence "t \<in> ?A" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def)
} ultimately show "?A = ?B" by auto
qed
lemma dual_transaction_ik_is_transaction_send':
fixes \<delta>::"('a,'b,'c) prot_subst"
assumes "wellformed_transaction T"
shows "ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>))) =
trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T)) \<cdot>\<^sub>s\<^sub>e\<^sub>t \<delta>" (is "?A = ?B")
using dual_transaction_ik_is_transaction_send[OF assms]
subst_lsst_unlabel[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" \<delta>]
ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" \<delta>]
dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" \<delta>]
by auto
lemma db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq:
assumes T: "wellformed_transaction T"
and S: "prefix S (transaction_receive T@transaction_selects T@transaction_checks T)"
shows "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>))"
proof -
let ?T1 = "transaction_receive T"
let ?T2 = "transaction_selects T"
let ?T3 = "transaction_checks T"
have *: "prefix (unlabel S) (unlabel (?T1@?T2@?T3))" using S prefix_proj(1) by blast
have "list_all is_Receive (unlabel ?T1)"
"list_all is_Assignment (unlabel ?T2)"
"list_all is_Check (unlabel ?T3)"
using T by (simp_all add: wellformed_transaction_def)
hence "\<forall>b \<in> set (unlabel ?T1). \<not>is_Insert b \<and> \<not>is_Delete b"
"\<forall>b \<in> set (unlabel ?T2). \<not>is_Insert b \<and> \<not>is_Delete b"
"\<forall>b \<in> set (unlabel ?T3). \<not>is_Insert b \<and> \<not>is_Delete b"
by (metis (mono_tags, lifting) Ball_set stateful_strand_step.distinct_disc(16,18),
metis (mono_tags, lifting) Ball_set stateful_strand_step.distinct_disc(24,26,33,37),
metis (mono_tags, lifting) Ball_set stateful_strand_step.distinct_disc(24,26,33,35,37,39))
hence "\<forall>b \<in> set (unlabel (?T1@?T2@?T3)). \<not>is_Insert b \<and> \<not>is_Delete b"
by (auto simp add: unlabel_def)
hence "\<forall>b \<in> set (unlabel S). \<not>is_Insert b \<and> \<not>is_Delete b"
using * unfolding prefix_def by fastforce
hence "\<forall>b \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \<cdot>\<^sub>s\<^sub>s\<^sub>t \<delta>). \<not>is_Insert b \<and> \<not>is_Delete b"
proof (induction S)
case (Cons a S)
then obtain l b where "a = (l,b)" by (metis surj_pair)
thus ?case
using Cons unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def unlabel_def subst_apply_stateful_strand_def
by (cases b) auto
qed simp
hence **: "\<forall>b \<in> set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<delta>))). \<not>is_Insert b \<and> \<not>is_Delete b"
by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel)
show ?thesis
using db\<^sub>s\<^sub>s\<^sub>t_no_upd_append[OF **] unlabel_append
unfolding db\<^sub>s\<^sub>s\<^sub>t_def by metis
qed
lemma db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_ex:
assumes "d \<in> set (db'\<^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 \<theta>) \<I> D)"
"\<forall>t u. insert\<langle>t,u\<rangle> \<in> set (unlabel A) \<longrightarrow> (\<exists>s. u = Fun (Set s) [])"
"\<forall>t u. delete\<langle>t,u\<rangle> \<in> set (unlabel A) \<longrightarrow> (\<exists>s. u = Fun (Set s) [])"
"\<forall>d \<in> set D. \<exists>s. snd d = Fun (Set s) []"
shows "\<exists>s. snd d = Fun (Set s) []"
using assms
proof (induction A arbitrary: D)
case (Cons a A)
obtain l b where a: "a = (l,b)" by (metis surj_pair)
have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = receive\<langle>t \<cdot> \<theta>\<rangle>#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
when "b = send\<langle>t\<rangle>" for t
by (simp add: a that subst_lsst_unlabel_cons)
have 2: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = send\<langle>t \<cdot> \<theta>\<rangle>#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
when "b = receive\<langle>t\<rangle>" for t
by (simp add: a that subst_lsst_unlabel_cons)
have 3: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>) = (b \<cdot>\<^sub>s\<^sub>s\<^sub>t\<^sub>p \<theta>)#unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \<cdot>\<^sub>l\<^sub>s\<^sub>s\<^sub>t \<theta>)"
when "\<nexists>t. b = send\<langle>t\<rangle> \<or> b = receive\<langle>t\<rangle>"
using a that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons subst_lsst_unlabel_cons[of l b]
by (cases b) auto
show ?case using 1 2 3 a Cons by (cases b) fastforce+
qed simp
lemma is_Fun_SetE[elim]:
assumes t: "is_Fun_Set t"
obtains s where "t = Fun (Set s) []"
proof (cases t)
case (Fun f T)
then obtain s where "f = Set s" using t unfolding is_Fun_Set_def by (cases f) moura+
moreover have "T = []" using Fun t unfolding is_Fun_Set_def by (cases T) auto
ultimately show ?thesis using Fun that by fast
qed (use t is_Fun_Set_def in fast)
lemma Fun_Set_InSet_iff:
"(u = \<langle>a: Var x \<in> Fun (Set s) []\<rangle>) \<longleftrightarrow>
(is_InSet u \<and> is_Var (the_elem_term u) \<and> is_Fun_Set (the_set_term u) \<and>
the_Set (the_Fun (the_set_term u)) = s \<and> the_Var (the_elem_term u) = x \<and> the_check u = a)"
(is "?A \<longleftrightarrow> ?B")
proof
show "?A \<Longrightarrow> ?B" unfolding is_Fun_Set_def by auto
assume B: ?B
thus ?A
proof (cases u)
case (InSet b t t')
hence "b = a" "t = Var x" "t' = Fun (Set s) []"
using B by (simp, fastforce, fastforce)
thus ?thesis using InSet by fast
qed auto
qed
lemma Fun_Set_NotInSet_iff:
"(u = \<langle>Var x not in Fun (Set s) []\<rangle>) \<longleftrightarrow>
(is_NegChecks u \<and> bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \<and> the_eqs u = [] \<and> length (the_ins u) = 1 \<and>
is_Var (fst (hd (the_ins u))) \<and> is_Fun_Set (snd (hd (the_ins u)))) \<and>
the_Set (the_Fun (snd (hd (the_ins u)))) = s \<and> the_Var (fst (hd (the_ins u))) = x"
(is "?A \<longleftrightarrow> ?B")
proof
show "?A \<Longrightarrow> ?B" unfolding is_Fun_Set_def by auto
assume B: ?B
show ?A
proof (cases u)
case (NegChecks X F F')
hence "X = []" "F = []"
using B by auto
moreover have "fst (hd (the_ins u)) = Var x" "snd (hd (the_ins u)) = Fun (Set s) []"
using B is_Fun_SetE[of "snd (hd (the_ins u))"]
by (force, fastforce)
hence "F' = [(Var x, Fun (Set s) [])]"
using NegChecks B by (cases "the_ins u") auto
ultimately show ?thesis using NegChecks by fast
qed (use B in auto)
qed
lemma is_Fun_Set_exi: "is_Fun_Set x \<longleftrightarrow> (\<exists>s. x = Fun (Set s) [])"
by (metis prot_fun.collapse(2) term.collapse(2) prot_fun.disc(15) term.disc(2)
term.sel(2,4) is_Fun_Set_def un_Fun1_def)
lemma is_Fun_Set_subst:
assumes "is_Fun_Set S'"
shows "is_Fun_Set (S' \<cdot> \<sigma>)"
using assms by (fastforce simp add: is_Fun_Set_def)
lemma is_Update_in_transaction_updates:
assumes tu: "is_Update t"
assumes t: "t \<in> set (unlabel (transaction_strand TT))"
assumes vt: "wellformed_transaction TT"
shows "t \<in> set (unlabel (transaction_updates TT))"
using t tu vt unfolding transaction_strand_def wellformed_transaction_def list_all_iff
by (auto simp add: unlabel_append)
lemma transaction_fresh_vars_subset:
assumes "wellformed_transaction T"
shows "set (transaction_fresh T) \<subseteq> fv_transaction T"
using assms fv_transaction_unfold[of T]
unfolding wellformed_transaction_def
by auto
lemma transaction_fresh_vars_notin:
assumes T: "wellformed_transaction T"
and x: "x \<in> set (transaction_fresh T)"
shows "x \<notin> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?A)
and "x \<notin> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" (is ?B)
and "x \<notin> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?C)
and "x \<notin> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?D)
and "x \<notin> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" (is ?E)
and "x \<notin> vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?F)
and "x \<notin> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?G)
and "x \<notin> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" (is ?H)
and "x \<notin> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?I)
proof -
have 0:
"set (transaction_fresh T) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"
"set (transaction_fresh T) \<inter> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}"
"set (transaction_fresh T) \<inter> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = {}"
"fv_transaction T \<inter> bvars_transaction T = {}"
"fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \<subseteq> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \<union> fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)"
using T unfolding wellformed_transaction_def
by fast+
have 1: "set (transaction_fresh T) \<inter> bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}"
using 0(1,4) fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast
have 2:
"vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)"
"vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)"
"bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}"
"bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = {}"
using bvars_wellformed_transaction_unfold[OF T] bvars_transaction_unfold[of T]
vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_receive T)"]
vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_selects T)"]
by blast+
show ?A ?B ?C ?D ?E ?G ?H ?I using 0 1 2 x by fast+
show ?F using 0(2,3,5) 1 x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_checks T)"] by fast
qed
lemma transaction_proj_member:
assumes "T \<in> set P"
shows "transaction_proj n T \<in> set (map (transaction_proj n) P)"
using assms by simp
lemma transaction_strand_proj:
"transaction_strand (transaction_proj n T) = proj n (transaction_strand T)"
proof -
obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp
thus ?thesis
using transaction_proj.simps[of n A B C D E F]
unfolding transaction_strand_def proj_def Let_def by auto
qed
lemma transaction_proj_fresh_eq:
"transaction_fresh (transaction_proj n T) = transaction_fresh T"
proof -
obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp
thus ?thesis
using transaction_proj.simps[of n A B C D E F]
unfolding transaction_fresh_def proj_def Let_def by auto
qed
lemma transaction_proj_trms_subset:
"trms_transaction (transaction_proj n T) \<subseteq> trms_transaction T"
proof -
obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp
thus ?thesis
using transaction_proj.simps[of n A B C D E F] trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of n]
unfolding transaction_fresh_def Let_def transaction_strand_def by auto
qed
lemma transaction_proj_vars_subset:
"vars_transaction (transaction_proj n T) \<subseteq> vars_transaction T"
proof -
obtain A B C D E F where "T = Transaction A B C D E F" by (cases T) simp
thus ?thesis
using transaction_proj.simps[of n A B C D E F]
sst_vars_proj_subset(3)[of n "transaction_strand T"]
unfolding transaction_fresh_def Let_def transaction_strand_def by simp
qed
end

View File

@ -0,0 +1,72 @@
@InProceedings{ brucker.ea:integrating:2009,
author = {Achim D. Brucker and Sebastian M{\"{o}}dersheim},
editor = {Pierpaolo Degano and Joshua D. Guttman},
title = {{Integrating Automated and Interactive Protocol Verification}},
booktitle = {Formal Aspects in Security and Trust, 6th International Workshop, {FAST} 2009, Eindhoven, The
Netherlands, November 5-6, 2009, Revised Selected Papers},
series = {Lecture Notes in Computer Science},
volume = 5983,
pages = {248--262},
publisher = {Springer},
year = 2009,
doi = {10.1007/978-3-642-12459-4_18}
}
@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}
}
@Article{ hess.ea:stateful:2020,
author = {Andreas V. Hess and Sebastian~M{\"o}dersheim and Achim~D.~Brucker},
title = {{Stateful Protocol Composition and Typing}},
journal = {Archive of Formal Proofs},
month = apr,
year = 2020,
note = {\url{http://isa-afp.org/entries/Stateful_Protocol_Composition_and_Typing.html}, Formal proof
development},
issn = {2150-914x}
}

View File

@ -0,0 +1,166 @@
\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}
\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{Automated Stateful Protocol Verification}
\author{%
\begin{minipage}{.8\textwidth}
\centering
\href{https://www.dtu.dk/english/service/phonebook/person?id=64207}{Andreas~V.~Hess}\footnotemark[1]
\qquad\qquad
\href{https://people.compute.dtu.dk/samo/}{Sebastian~M{\"o}dersheim}\footnotemark[1]
\\
\href{http://www.brucker.ch/}{Achim~D.~Brucker}\footnotemark[2]
\qquad\qquad
\href{https://people.compute.dtu.dk/andschl}{Anders~Schlichtkrull}
\end{minipage}
}
\publishers{%
\footnotemark[1]~DTU Compute, Technical University of Denmark, Lyngby, Denmark\texorpdfstring{\\}{, }
\texttt{\{avhe, samo, andschl\}@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}
In protocol verification we observe a wide spectrum from fully
automated methods to interactive theorem proving with proof
assistants like Isabelle/HOL.
In this AFP entry, we present a fully-automated approach for
verifying stateful security protocols, i.e., protocols with mutable
state that may span several sessions.
The approach supports reachability goals like secrecy and
authentication.
We also include a simple user-friendly transaction-based
protocol specification language that is embedded into Isabelle.
\bigskip
\noindent{\textbf{Keywords:}}
Fully automated verification, stateful security protocols
\end{quote}
\end{abstract}
\tableofcontents
\cleardoublepage
\chapter{Introduction}
In protocol verification we observe a wide spectrum from fully
automated methods to interactive theorem proving with proof
assistants like Isabelle/HOL. The latter provide overwhelmingly high
assurance of the correctness, which automated methods often cannot:
due to their complexity, bugs in such automated verification tools
are likely and thus the risk of erroneously verifying a flawed
protocol is non-negligible. There are a few works that try to
combine advantages from both ends of the spectrum: a high degree of
automation and assurance.
Inspired by~\cite{brucker.ea:integrating:2009}, we present here a
first step towards achieving this for a more challenging class of
protocols, namely those that work with a mutable long-term state. To
our knowledge this is the first approach that achieves fully
automated verification of stateful protocols in an LCF-style theorem
prover. The approach also includes a simple user-friendly
transaction-based protocol specification language embedded into
Isabelle, and can also leverage a number of existing results such as
soundness of a typed model (see,
e.g.,~\cite{hess:typing:2018,hess.ea:formalizing:2017,hess.ea:typing:2018})
and compositionality (see,
e.g.,~\cite{hess:typing:2018,hess.ea:stateful:2018}). The Isabelle
formalization extends the AFP entry on stateful protocol composition and
typing~\cite{hess.ea:stateful:2020}.
\begin{figure}
\centering
\includegraphics[height=\textheight]{session_graph}
\caption{The Dependency Graph of the Isabelle Theories.\label{fig:session-graph}}
\end{figure}
The rest of this document is automatically generated from the
formalization in Isabelle/HOL, i.e., all content is checked by
Isabelle. Overall, the structure of this document follows the
theory dependencies (see \autoref{fig:session-graph}): We start with
the formal framework for verifying stateful security protocols
(\autoref{cha:verification}). We continue with the setup for
supporting the high-level protocol specifications language for
security protocols (the Trac format) and the implementation of the
fully automated proof tactics (\autoref{cha:trac}). Finally, we
present examples (\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
\chapter{Stateful Protocol Verification}
\label{cha:verification}
\input{Transactions.tex}
\input{Term_Abstraction.tex}
\input{Stateful_Protocol_Model.tex}
\input{Term_Variants.tex}
\input{Term_Implication.tex}
\input{Stateful_Protocol_Verification.tex}
\chapter{Trac Support and Automation}
\label{cha:trac}
\input{Eisbach_Protocol_Verification.tex}
\input{ml_yacc_lib.tex}
\input{trac_term.tex}
\input{trac_fp_parser.tex}
\input{trac_protocol_parser.tex}
\input{trac.tex}
\chapter{Examples}
\label{cha:examples}
\input{Keyserver.tex}
\input{Keyserver2.tex}
\input{Keyserver_Composition.tex}
\input{PKCS_Model03.tex}
\input{PKCS_Model07.tex}
\input{PKCS_Model09.tex}
% \input{session}
{\small
\bibliographystyle{abbrvnat}
\bibliography{root}
}
\end{document}
\endinput
%%% Local Variables:
%%% mode: latex
%%% TeX-master: t
%%% End:

View File

@ -0,0 +1,133 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Keyserver.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>The Keyserver Protocol\<close>
theory Keyserver
imports "../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: keyserver
Types:
honest = {a,b,c}
server = {s}
agents = honest ++ server
Sets:
ring/1 valid/2 revoked/2
Functions:
Public sign/2 crypt/2 pair/2
Private inv/1
Analysis:
sign(X,Y) -> Y
crypt(X,Y) ? inv(X) -> Y
pair(X,Y) -> X,Y
Transactions:
# Out-of-band registration
outOfBand(A:honest,S:server)
new NPK
insert NPK ring(A)
insert NPK valid(A,S)
send NPK.
# User update key
keyUpdateUser(A:honest,PK:value)
PK in ring(A)
new NPK
delete PK ring(A)
insert NPK ring(A)
send sign(inv(PK),pair(A,NPK)).
# Server update key
keyUpdateServer(A:honest,S:server,PK:value,NPK:value)
receive sign(inv(PK),pair(A,NPK))
PK in valid(A,S)
NPK notin valid(_)
NPK notin revoked(_)
delete PK valid(A,S)
insert PK revoked(A,S)
insert NPK valid(A,S)
send inv(PK).
# Attack definition
authAttack(A:honest,S:server,PK:value)
receive inv(PK)
PK in valid(A,S)
attack.
\<close>\<open>
val(ring(A)) where A:honest
sign(inv(val(0)),pair(A,val(ring(A)))) where A:honest
inv(val(revoked(A,S))) where A:honest S:server
pair(A,val(ring(A))) where A:honest
occurs(val(ring(A))) where A:honest
timplies(val(ring(A)),val(ring(A),valid(A,S))) where A:honest S:server
timplies(val(ring(A)),val(0)) where A:honest
timplies(val(ring(A),valid(A,S)),val(valid(A,S))) where A:honest S:server
timplies(val(0),val(valid(A,S))) where A:honest S:server
timplies(val(valid(A,S)),val(revoked(A,S))) where A:honest S:server
\<close>
subsection \<open>Proof of security\<close>
protocol_model_setup spm: keyserver
compute_SMP [optimized] keyserver_protocol keyserver_SMP
manual_protocol_security_proof ssp: keyserver
for keyserver_protocol keyserver_fixpoint keyserver_SMP
apply check_protocol_intro
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
done
end

View File

@ -0,0 +1,132 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Keyserver2.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>A Variant of the Keyserver Protocol\<close>
theory Keyserver2
imports "../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: keyserver2
Types:
honest = {a,b,c}
dishonest = {i}
agent = honest ++ dishonest
Sets:
ring'/1 seen/1 pubkeys/0 valid/1
Functions:
Public h/1 sign/2 crypt/2 scrypt/2 pair/2 update/3
Private inv/1 pw/1
Analysis:
sign(X,Y) -> Y
crypt(X,Y) ? inv(X) -> Y
scrypt(X,Y) ? X -> Y
pair(X,Y) -> X,Y
update(X,Y,Z) -> X,Y,Z
Transactions:
passwordGenD(A:dishonest)
send pw(A).
pubkeysGen()
new PK
insert PK pubkeys
send PK.
updateKeyPw(A:honest,PK:value)
PK in pubkeys
new NPK
insert NPK ring'(A)
send NPK
send crypt(PK,update(A,NPK,pw(A))).
updateKeyServerPw(A:agent,PK:value,NPK:value)
receive crypt(PK,update(A,NPK,pw(A)))
PK in pubkeys
NPK notin pubkeys
NPK notin seen(_)
insert NPK valid(A)
insert NPK seen(A).
authAttack2(A:honest,PK:value)
receive inv(PK)
PK in valid(A)
attack.
\<close>
subsection \<open>Proof of security \<close>
protocol_model_setup spm: keyserver2
compute_fixpoint keyserver2_protocol keyserver2_fixpoint
protocol_security_proof ssp: keyserver2
subsection \<open>The generated theorems and definitions\<close>
thm ssp.protocol_secure
thm keyserver2_enum_consts.nchotomy
thm keyserver2_sets.nchotomy
thm keyserver2_fun.nchotomy
thm keyserver2_atom.nchotomy
thm keyserver2_arity.simps
thm keyserver2_public.simps
thm keyserver2_\<Gamma>.simps
thm keyserver2_Ana.simps
thm keyserver2_transaction_passwordGenD_def
thm keyserver2_transaction_pubkeysGen_def
thm keyserver2_transaction_updateKeyPw_def
thm keyserver2_transaction_updateKeyServerPw_def
thm keyserver2_transaction_authAttack2_def
thm keyserver2_protocol_def
thm keyserver2_fixpoint_def
end

View File

@ -0,0 +1,295 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: Keyserver_Composition.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>The Composition of the Two Keyserver Protocols\<close>
theory Keyserver_Composition
imports "../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: kscomp
Types:
honest = {a,b,c}
dishonest = {i}
agent = honest ++ dishonest
Sets:
ring/1 valid/1 revoked/1 deleted/1
ring'/1 seen/1 pubkeys/0
Functions:
Public h/1 sign/2 crypt/2 scrypt/2 pair/2 update/3
Private inv/1 pw/1
Analysis:
sign(X,Y) -> Y
crypt(X,Y) ? inv(X) -> Y
scrypt(X,Y) ? X -> Y
pair(X,Y) -> X,Y
update(X,Y,Z) -> X,Y,Z
Transactions:
### The signature-based keyserver protocol
p1_outOfBand(A:honest)
new PK
insert PK ring(A)
* insert PK valid(A)
send PK.
p1_oufOfBandD(A:dishonest)
new PK
* insert PK valid(A)
send PK
send inv(PK).
p1_updateKey(A:honest,PK:value)
PK in ring(A)
new NPK
delete PK ring(A)
insert PK deleted(A)
insert NPK ring(A)
send sign(inv(PK),pair(A,NPK)).
p1_updateKeyServer(A:agent,PK:value,NPK:value)
receive sign(inv(PK),pair(A,NPK))
* PK in valid(A)
* NPK notin valid(_)
NPK notin revoked(_)
* delete PK valid(A)
insert PK revoked(A)
* insert NPK valid(A)
send inv(PK).
p1_authAttack(A:honest,PK:value)
receive inv(PK)
* PK in valid(A)
attack.
### The password-based keyserver protocol
p2_passwordGenD(A:dishonest)
send pw(A).
p2_pubkeysGen()
new PK
insert PK pubkeys
send PK.
p2_updateKeyPw(A:honest,PK:value)
PK in pubkeys
new NPK
# NOTE: The ring' sets are not used elsewhere, but we have to avoid that the fresh keys generated
# by this rule are abstracted to the empty abstraction, and so we insert them into a ring'
# set. Otherwise the two protocols would have too many abstractions in common (in particular,
# the empty abstraction) which leads to false attacks in the composed protocol (probably
# because the term implication graphs of the two protocols then become 'linked' through the
# empty abstraction)
insert NPK ring'(A)
send NPK
send crypt(PK,update(A,NPK,pw(A))).
#Transactions of p2:
p2_updateKeyServerPw(A:agent,PK:value,NPK:value)
receive crypt(PK,update(A,NPK,pw(A)))
PK in pubkeys
NPK notin pubkeys
NPK notin seen(_)
* insert NPK valid(A)
insert NPK seen(A).
p2_authAttack2(A:honest,PK:value)
receive inv(PK)
* PK in valid(A)
attack.
\<close> \<open>
sign(inv(val(deleted(A))),pair(A,val(ring(A)))) where A:honest
sign(inv(val(deleted(A),valid(B))),pair(A,val(ring(A)))) where A:honest B:dishonest
sign(inv(val(deleted(A),seen(B),valid(B))),pair(A,val(ring(A)))) where A:honest B:dishonest
sign(inv(val(deleted(A),valid(A))),pair(A,val(ring(A)))) where A:honest B:dishonest
sign(inv(val(deleted(A),seen(B),valid(B),valid(A))),pair(A,val(ring(A)))) where A:honest B:dishonest
pair(A,val(ring(A))) where A:honest
inv(val(deleted(A),revoked(A))) where A:honest
inv(val(valid(A))) where A:dishonest
inv(val(revoked(A))) where A:dishonest
inv(val(revoked(A),seen(A))) where A:dishonest
inv(val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
inv(val(revoked(A),deleted(A),seen(B),valid(B))) where A:honest B:dishonest
occurs(val(ring(A))) where A:honest
occurs(val(valid(A))) where A:dishonest
occurs(val(ring'(A))) where A:honest
occurs(val(pubkeys))
occurs(val(valid(A),ring(A))) where A:honest
pw(A) where A:dishonest
crypt(val(pubkeys),update(A,val(ring'(A)),pw(A))) where A:honest
val(ring(A)) where A:honest
val(valid(A)) where A:dishonest
val(ring'(A)) where A:honest
val(pubkeys)
val(valid(A),ring(A)) where A:honest
timplies(val(pubkeys),val(valid(A),pubkeys)) where A:dishonest
timplies(val(ring'(A)),val(ring'(A),valid(B))) where A:honest B:dishonest
timplies(val(ring'(A)),val(ring'(A),valid(A),seen(A))) where A:honest
timplies(val(ring'(A)),val(ring'(A),valid(A),seen(A),valid(B))) where A:honest B:dishonest
timplies(val(ring'(A)),val(seen(B),valid(B),ring'(A))) where A:honest B:dishonest
timplies(val(ring'(A),valid(B)),val(ring'(A),valid(A),seen(A),valid(B))) where A:honest B:dishonest
timplies(val(ring'(A),valid(B)),val(seen(B),valid(B),ring'(A))) where A:honest B:dishonest
timplies(val(ring(A)),val(ring(A),valid(A))) where A:honest
timplies(val(ring(A)),val(ring(A),valid(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(deleted(A))) where A:honest
timplies(val(ring(A)),val(revoked(A),deleted(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(revoked(A),deleted(A),seen(B),revoked(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(deleted(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(ring(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(valid(A),deleted(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A)),val(valid(A),ring(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A),valid(A)),val(deleted(A),valid(A))) where A:honest
timplies(val(ring(A),valid(B)),val(deleted(A),valid(B))) where A:honest B:dishonest
timplies(val(ring(A),valid(A)),val(deleted(A),revoked(A))) where A:honest
timplies(val(deleted(A)),val(deleted(A),valid(A))) where A:honest
timplies(val(deleted(A)),val(deleted(A),valid(B))) where A:honest B:dishonest
timplies(val(deleted(A)),val(revoked(A),seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(deleted(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(deleted(A)),val(seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(deleted(A)),val(seen(B),valid(B),valid(A),deleted(A))) where A:honest B:dishonest
timplies(val(revoked(A)),val(seen(A),revoked(A))) where A:dishonest
timplies(val(revoked(A)),val(seen(A),revoked(A),valid(A))) where A:dishonest
timplies(val(revoked(A),deleted(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(revoked(A),deleted(A)),val(seen(B),valid(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),deleted(A),valid(A)),val(revoked(A),seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),deleted(A),valid(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),revoked(A),deleted(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(seen(A),valid(A)),val(revoked(A),seen(A))) where A:dishonest
timplies(val(seen(A),valid(A),revoked(A)),val(seen(A),revoked(A))) where A:dishonest
timplies(val(seen(B),valid(B),ring(A)),val(deleted(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),valid(A),ring(A)),val(deleted(A),seen(B),valid(B),valid(A))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),valid(A),ring(A)),val(revoked(A),seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(seen(B),valid(B),valid(A),ring(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(valid(A)),val(revoked(A))) where A:dishonest
timplies(val(valid(A),deleted(A)),val(deleted(A),revoked(A))) where A:honest
timplies(val(valid(A),deleted(A)),val(revoked(A),seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(valid(A),deleted(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(valid(A),deleted(A)),val(seen(B),valid(B),valid(A),deleted(A))) where A:honest B:dishonest
timplies(val(ring(A),valid(A)),val(deleted(A),seen(B),valid(B),valid(A))) where A:honest B:dishonest
timplies(val(ring(A),valid(A)),val(revoked(B),seen(B),revoked(A),deleted(A))) where A:honest B:dishonest
timplies(val(ring(A),valid(A)),val(seen(B),valid(B),valid(A),ring(A))) where A:honest B:dishonest
timplies(val(valid(B),deleted(A)),val(seen(B),valid(B),deleted(A))) where A:honest B:dishonest
timplies(val(ring(A),valid(B)),val(deleted(A),seen(B),valid(B))) where A:honest B:dishonest
timplies(val(ring(A),valid(B)),val(seen(B),valid(B),ring(A))) where A:honest B:dishonest
timplies(val(valid(A)),val(seen(A),valid(A))) where A:dishonest
\<close>
subsection \<open>Proof: The composition of the two keyserver protocols is secure\<close>
protocol_model_setup spm: kscomp
setup_protocol_checks spm kscomp_protocol
manual_protocol_security_proof ssp: kscomp
apply check_protocol_intro
subgoal by code_simp
subgoal
apply coverage_check_intro
subgoal by code_simp
subgoal by code_simp
subgoal by eval
subgoal by eval
subgoal by eval
subgoal by code_simp
subgoal by code_simp
subgoal by eval
subgoal by eval
subgoal by eval
done
subgoal by eval
subgoal by eval
subgoal
apply (unfold spm.wellformed_fixpoint_def Let_def case_prod_unfold; intro conjI)
subgoal by code_simp
subgoal by code_simp
subgoal by eval
subgoal by code_simp
subgoal by code_simp
done
done
subsection \<open>The generated theorems and definitions\<close>
thm ssp.protocol_secure
thm kscomp_enum_consts.nchotomy
thm kscomp_sets.nchotomy
thm kscomp_fun.nchotomy
thm kscomp_atom.nchotomy
thm kscomp_arity.simps
thm kscomp_public.simps
thm kscomp_\<Gamma>.simps
thm kscomp_Ana.simps
thm kscomp_transaction_p1_outOfBand_def
thm kscomp_transaction_p1_oufOfBandD_def
thm kscomp_transaction_p1_updateKey_def
thm kscomp_transaction_p1_updateKeyServer_def
thm kscomp_transaction_p1_authAttack_def
thm kscomp_transaction_p2_passwordGenD_def
thm kscomp_transaction_p2_pubkeysGen_def
thm kscomp_transaction_p2_updateKeyPw_def
thm kscomp_transaction_p2_updateKeyServerPw_def
thm kscomp_transaction_p2_authAttack2_def
thm kscomp_protocol_def
thm kscomp_fixpoint_def
end

View File

@ -0,0 +1,161 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: PKCS_Model03.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>The PKCS Model, Scenario 3\<close>
theory PKCS_Model03
imports "../../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: ATTACK_UNSET
Types:
token = {token1}
Sets:
extract/1 wrap/1 decrypt/1 sensitive/1
Functions:
Public senc/2 h/1
Private inv/1
Analysis:
senc(M,K2) ? K2 -> M #This analysis rule corresponds to the decrypt2 rule in the AIF-omega specification.
#M was type untyped
Transactions:
iik1()
new K1
insert K1 sensitive(token1)
insert K1 extract(token1)
send h(K1).
iik2()
new K2
insert K2 wrap(token1)
send h(K2).
# ======================wrap================
wrap(K1:value,K2:value)
receive h(K1)
receive h(K2)
K1 in extract(token1)
K2 in wrap(token1)
send senc(K1,K2).
# ======================set wrap================
setwrap(K2:value)
receive h(K2)
K2 notin decrypt(token1)
insert K2 wrap(token1).
# ======================set decrypt================
setdecrypt(K2:value)
receive h(K2)
K2 notin wrap(token1)
insert K2 decrypt(token1).
# ======================decrypt================
decrypt1(K2:value,M:value) #M was untyped in the AIF-omega specification.
receive h(K2)
receive senc(M,K2)
K2 in decrypt(token1)
send M.
# ======================attacks================
attack1(K1:value)
receive K1
K1 in sensitive(token1)
attack.
\<close>
subsection \<open>Protocol model setup\<close>
protocol_model_setup spm: ATTACK_UNSET
subsection \<open>Fixpoint computation\<close>
compute_fixpoint ATTACK_UNSET_protocol ATTACK_UNSET_fixpoint
compute_SMP [optimized] ATTACK_UNSET_protocol ATTACK_UNSET_SMP
subsection \<open>Proof of security\<close>
manual_protocol_security_proof ssp: ATTACK_UNSET
for ATTACK_UNSET_protocol ATTACK_UNSET_fixpoint ATTACK_UNSET_SMP
apply check_protocol_intro
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
subgoal by code_simp
done
subsection \<open>The generated theorems and definitions\<close>
thm ssp.protocol_secure
thm ATTACK_UNSET_enum_consts.nchotomy
thm ATTACK_UNSET_sets.nchotomy
thm ATTACK_UNSET_fun.nchotomy
thm ATTACK_UNSET_atom.nchotomy
thm ATTACK_UNSET_arity.simps
thm ATTACK_UNSET_public.simps
thm ATTACK_UNSET_\<Gamma>.simps
thm ATTACK_UNSET_Ana.simps
thm ATTACK_UNSET_transaction_iik1_def
thm ATTACK_UNSET_transaction_iik2_def
thm ATTACK_UNSET_transaction_wrap_def
thm ATTACK_UNSET_transaction_setwrap_def
thm ATTACK_UNSET_transaction_setdecrypt_def
thm ATTACK_UNSET_transaction_decrypt1_def
thm ATTACK_UNSET_transaction_attack1_def
thm ATTACK_UNSET_protocol_def
thm ATTACK_UNSET_fixpoint_def
thm ATTACK_UNSET_SMP_def
end

View File

@ -0,0 +1,221 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: PKCS_Model07.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>The PKCS Protocol, Scenario 7\<close>
theory PKCS_Model07
imports "../../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: RE_IMPORT_ATT
Types:
token = {token1}
Sets:
extract/1 wrap/1 unwrap/1 decrypt/1 sensitive/1
Functions:
Public senc/2 h/2 bind/2
Private inv/1
Analysis:
senc(M1,K2) ? K2 -> M1 #This analysis rule corresponds to the decrypt2 rule in the AIF-omega specification.
#M1 was type untyped
Transactions:
iik1()
new K1
new N1
insert N1 sensitive(token1)
insert N1 extract(token1)
insert K1 sensitive(token1)
send h(N1,K1).
iik2()
new K2
new N2
insert N2 wrap(token1)
insert N2 extract(token1)
send h(N2,K2).
# =====set wrap=====
setwrap(N2:value,K2:value)
receive h(N2,K2)
N2 notin sensitive(token1)
N2 notin decrypt(token1)
insert N2 wrap(token1).
# =====set unwrap===
setunwrap(N2:value,K2:value)
receive h(N2,K2)
N2 notin sensitive(token1)
insert N2 unwrap(token1).
# =====unwrap, generate new handler======
#-----------the senstive attr copy-------------
unwrapsensitive(M2:value, K2:value, N1:value, N2:value) #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2)
receive h(N2,K2)
N1 in sensitive(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew sensitive(token1)
send h(Nnew,M2).
#-----------the wrap attr copy-------------
wrapattr(M2:value, K2:value, N1:value, N2:value) #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2)
receive h(N2,K2)
N1 in wrap(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew wrap(token1)
send h(Nnew,M2).
#-----------the decrypt attr copy-------------
decrypt1attr(M2:value,K2:value,N1:value,N2:value) #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2)
receive h(N2,K2)
N1 in decrypt(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew decrypt(token1)
send h(Nnew,M2).
decrypt2attr(M2:value,K2:value,N1:value,N2:value) #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2)
receive h(N2,K2)
N1 notin sensitive(token1)
N1 notin wrap(token1)
N1 notin decrypt(token1)
N2 in unwrap(token1)
new Nnew
send h(Nnew,M2).
# ======================wrap================
wrap(N1:value,K1:value,N2:value,K2:value)
receive h(N1,K1)
receive h(N2,K2)
N1 in extract(token1)
N2 in wrap(token1)
send senc(K1,K2)
send bind(N1,K1).
# =====set decrypt===
setdecrypt(Nnew:value, K2:value)
receive h(Nnew,K2)
Nnew notin wrap(token1)
insert Nnew decrypt(token1).
# ======================decrypt================
decrypt1(Nnew:value, K2:value,M1:value) #M1 was untyped in the AIF-omega specification.
receive h(Nnew,K2)
receive senc(M1,K2)
Nnew in decrypt(token1)
delete Nnew decrypt(token1)
send M1.
# ======================attacks================
attack1(K1:value)
receive K1
K1 in sensitive(token1)
attack.
\<close>
subsection \<open>Protocol model setup\<close>
protocol_model_setup spm: RE_IMPORT_ATT
subsection \<open>Fixpoint computation\<close>
compute_fixpoint RE_IMPORT_ATT_protocol RE_IMPORT_ATT_fixpoint
compute_SMP [optimized] RE_IMPORT_ATT_protocol RE_IMPORT_ATT_SMP
subsection \<open>Proof of security\<close>
protocol_security_proof [unsafe] ssp: RE_IMPORT_ATT
for RE_IMPORT_ATT_protocol RE_IMPORT_ATT_fixpoint RE_IMPORT_ATT_SMP
subsection \<open>The generated theorems and definitions\<close>
thm ssp.protocol_secure
thm RE_IMPORT_ATT_enum_consts.nchotomy
thm RE_IMPORT_ATT_sets.nchotomy
thm RE_IMPORT_ATT_fun.nchotomy
thm RE_IMPORT_ATT_atom.nchotomy
thm RE_IMPORT_ATT_arity.simps
thm RE_IMPORT_ATT_public.simps
thm RE_IMPORT_ATT_\<Gamma>.simps
thm RE_IMPORT_ATT_Ana.simps
thm RE_IMPORT_ATT_transaction_iik1_def
thm RE_IMPORT_ATT_transaction_iik2_def
thm RE_IMPORT_ATT_transaction_setwrap_def
thm RE_IMPORT_ATT_transaction_setunwrap_def
thm RE_IMPORT_ATT_transaction_unwrapsensitive_def
thm RE_IMPORT_ATT_transaction_wrapattr_def
thm RE_IMPORT_ATT_transaction_decrypt1attr_def
thm RE_IMPORT_ATT_transaction_decrypt2attr_def
thm RE_IMPORT_ATT_transaction_wrap_def
thm RE_IMPORT_ATT_transaction_setdecrypt_def
thm RE_IMPORT_ATT_transaction_decrypt1_def
thm RE_IMPORT_ATT_transaction_attack1_def
thm RE_IMPORT_ATT_protocol_def
thm RE_IMPORT_ATT_fixpoint_def
thm RE_IMPORT_ATT_SMP_def
end

View File

@ -0,0 +1,237 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: PKCS_Model09.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>The PKCS Protocol, Scenario 9\<close>
theory PKCS_Model09
imports "../../PSPSP"
begin
declare [[code_timing]]
trac\<open>
Protocol: LOSS_KEY_ATT
Types:
token = {token1}
Sets:
extract/1 wrap/1 unwrap/1 decrypt/1 sensitive/1
Functions:
Public senc/2 h/2 bind/3
Private inv/1
Analysis:
senc(M1,K2) ? K2 -> M1 #This analysis rule corresponds to the decrypt2 rule in the AIF-omega specification.
#M1 was type untyped
Transactions:
iik1()
new K1
new N1
insert N1 sensitive(token1)
insert N1 extract(token1)
insert K1 sensitive(token1)
send h(N1,K1).
iik2()
new K2
new N2
insert N2 wrap(token1)
insert N2 extract(token1)
send h(N2,K2).
iik3()
new K3
new N3
insert N3 extract(token1)
insert N3 decrypt(token1)
insert K3 decrypt(token1)
send h(N3,K3)
send K3.
# =====set wrap=====
setwrap(N2:value,K2:value) where N2 != K2
receive h(N2,K2)
N2 notin sensitive(token1)
N2 notin decrypt(token1)
insert N2 wrap(token1).
# =====set unwrap===
setunwrap(N2:value,K2:value) where N2 != K2
receive h(N2,K2)
N2 notin sensitive(token1)
insert N2 unwrap(token1).
# =====unwrap, generate new handler======
#-----------add the wrap attr copy-------------
unwrapWrap(M2:value,K2:value,N1:value,N2:value) where M2 != K2, M2 != N1, M2 != N2, K2 != N1, K2 != N2, N1 != N2 #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2,K2)
receive h(N2,K2)
N1 in wrap(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew wrap(token1)
send h(Nnew,M2).
#-----------add the senstive attr copy-------------
unwrapSens(M2:value,K2:value,N1:value,N2:value) where M2 != K2, M2 != N1, M2 != N2, K2 != N1, K2 != N2, N1 != N2 #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2,K2)
receive h(N2,K2)
N1 in sensitive(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew sensitive(token1)
send h(Nnew,M2).
#-----------add the decrypt attr copy-------------
decrypt1Attr(M2:value, K2:value,N1:value,N2:value) where M2 != K2, M2 != N1, M2 != N2, K2 != N1, K2 != N2, N1 != N2 #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2,K2)
receive h(N2,K2)
N1 in decrypt(token1)
N2 in unwrap(token1)
new Nnew
insert Nnew decrypt(token1)
send h(Nnew,M2).
decrypt2Attr(M2:value, K2:value,N1:value,N2:value) where M2 != K2, M2 != N1, M2 != N2, K2 != N1, K2 != N2, N1 != N2 #M2 was untyped in the AIF-omega specification.
receive senc(M2,K2)
receive bind(N1,M2,K2)
receive h(N2,K2)
N1 notin wrap(token1)
N1 notin sensitive(token1)
N1 notin decrypt(token1)
N2 in unwrap(token1)
new Nnew
send h(Nnew,M2).
# ======================wrap================
wrap(N1:value,K1:value, N2:value, K2:value) where N1 != N2, N1 != K2, N1 != K1, N2 != K2, N2 != K1, K2 != K1
receive h(N1,K1)
receive h(N2,K2)
N1 in extract(token1)
N2 in wrap(token1)
send senc(K1,K2)
send bind(N1,K1,K2).
# ======================bind generation================
bind1(K3:value,N2:value,K2:value, K1:value) where K3 != N2, K3 != K2, K3 != K1, N2 != K2, N2 != K1, K2 != K1
receive K3
receive h(N2,K2)
send bind(N2,K3,K3).
bind2(K3:value,N2:value,K2:value, K1:value) where K3 != N2, K3 != K2, K3 != K1, N2 != K2, N2 != K1, K2 != K1
receive K3
receive K1
receive h(N2,K2)
send bind(N2,K1,K3)
send bind(N2,K3,K1).
# =====set decrypt===
setdecrypt(Nnew:value,K2:value) where Nnew != K2
receive h(Nnew,K2)
Nnew notin wrap(token1)
insert Nnew decrypt(token1).
# ======================decrypt================
decrypt1(Nnew:value,K2:value,M1:value) where Nnew != K2, Nnew != M1, K2 != M1 #M1 was untyped in the AIF-omega specification.
receive h(Nnew,K2)
receive senc(M1,K2)
Nnew in decrypt(token1)
send M1.
# ======================attacks================
attack1(K1:value)
receive K1
K1 in sensitive(token1)
attack.
\<close>
subsection \<open>Protocol model setup\<close>
protocol_model_setup spm: LOSS_KEY_ATT
subsection \<open>Fixpoint computation\<close>
compute_fixpoint LOSS_KEY_ATT_protocol LOSS_KEY_ATT_fixpoint
text \<open>The fixpoint contains an attack signal\<close>
value "attack_notin_fixpoint LOSS_KEY_ATT_fixpoint"
subsection \<open>The generated theorems and definitions\<close>
thm LOSS_KEY_ATT_enum_consts.nchotomy
thm LOSS_KEY_ATT_sets.nchotomy
thm LOSS_KEY_ATT_fun.nchotomy
thm LOSS_KEY_ATT_atom.nchotomy
thm LOSS_KEY_ATT_arity.simps
thm LOSS_KEY_ATT_public.simps
thm LOSS_KEY_ATT_\<Gamma>.simps
thm LOSS_KEY_ATT_Ana.simps
thm LOSS_KEY_ATT_transaction_iik1_def
thm LOSS_KEY_ATT_transaction_iik2_def
thm LOSS_KEY_ATT_transaction_iik3_def
thm LOSS_KEY_ATT_transaction_setwrap_def
thm LOSS_KEY_ATT_transaction_setunwrap_def
thm LOSS_KEY_ATT_transaction_unwrapWrap_def
thm LOSS_KEY_ATT_transaction_unwrapSens_def
thm LOSS_KEY_ATT_transaction_decrypt1Attr_def
thm LOSS_KEY_ATT_transaction_decrypt2Attr_def
thm LOSS_KEY_ATT_transaction_wrap_def
thm LOSS_KEY_ATT_transaction_bind1_def
thm LOSS_KEY_ATT_transaction_bind2_def
thm LOSS_KEY_ATT_transaction_setdecrypt_def
thm LOSS_KEY_ATT_transaction_decrypt1_def
thm LOSS_KEY_ATT_transaction_attack1_def
thm LOSS_KEY_ATT_protocol_def
thm LOSS_KEY_ATT_fixpoint_def
end

View File

@ -0,0 +1,51 @@
#!/bin/sh
# (C) Copyright Andreas Viktor Hess, DTU, 2020
# (C) Copyright Sebastian A. Mödersheim, DTU, 2020
# (C) Copyright Achim D. Brucker, University of Exeter, 2020
# (C) Copyright Anders Schlichtkrull, DTU, 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.
ISABELLE=isabelle
all: trac_parser/trac_fp.lex.sml trac_parser/trac_fp.grm.sig trac_parser/trac_protocol.lex.sml trac_parser/trac_protocol.grm.sig
test:
isabelle build -c -D .
clean:
rm -f trac_parser/*.lex.sml trac_parser/*.grm.sml trac_parser/*.grm.sig
%.lex.sml: %.lex
bin/ml-lex-isa $<
%.grm.sig: %.grm
bin/ml-yacc-isa $<

View File

@ -0,0 +1,13 @@
# Interface between Isabelle and trac specifications
## Prerequisites
* For re-generating the parser, ml-lex and ml-yacc are required
## License
This project is licensed under a 2-clause BSD-style license.
SPDX-License-Identifier: BSD-2-Clause

View File

@ -0,0 +1,42 @@
#!/bin/bash
# (C) Copyright Andreas Viktor Hess, DTU, 2020
# (C) Copyright Sebastian A. Mödersheim, DTU, 2020
# (C) Copyright Achim D. Brucker, University of Exeter, 2020
# (C) Copyright Anders Schlichtkrull, DTU, 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.
ml-lex "$1"
sed -i -e '1s/^/ (***** GENERATED FILE -- DO NOT EDIT ****)\n/'\
-e "s/\\bref\\b/Unsynchronized.ref/g" \
-e "s/\\bUnsafe.\\b//g" \
-e "s/structure YYPosInt : INTEGER = Int/structure YYPosInt = Int/" \
"$1.sml"

View File

@ -0,0 +1,39 @@
#!/bin/bash
# (C) Copyright Andreas Viktor Hess, DTU, 2020
# (C) Copyright Sebastian A. Mödersheim, DTU, 2020
# (C) Copyright Achim D. Brucker, University of Exeter, 2020
# (C) Copyright Anders Schlichtkrull, DTU, 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.
ml-yacc "$1"
sed -i -e '1s/^/ (***** GENERATED FILE -- DO NOT EDIT ****)\n/'\
-e "s/\\bref\\b/Unsynchronized.ref/g" "$1.sml"

View File

@ -0,0 +1,323 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
(* base.sig: Base signature file for SML-Yacc. This file contains signatures
that must be loaded before any of the files produced by ML-Yacc are loaded
*)
(* STREAM: signature for a lazy stream.*)
signature STREAM =
sig type 'xa stream
val streamify : (unit -> '_a) -> '_a stream
val cons : '_a * '_a stream -> '_a stream
val get : '_a stream -> '_a * '_a stream
end
(* LR_TABLE: signature for an LR Table.
The list of actions and gotos passed to mkLrTable must be ordered by state
number. The values for state 0 are the first in the list, the values for
state 1 are next, etc.
*)
signature LR_TABLE =
sig
datatype ('a,'b) pairlist = EMPTY | PAIR of 'a * 'b * ('a,'b) pairlist
datatype state = STATE of int
datatype term = T of int
datatype nonterm = NT of int
datatype action = SHIFT of state
| REDUCE of int
| ACCEPT
| ERROR
type table
val numStates : table -> int
val numRules : table -> int
val describeActions : table -> state ->
(term,action) pairlist * action
val describeGoto : table -> state -> (nonterm,state) pairlist
val action : table -> state * term -> action
val goto : table -> state * nonterm -> state
val initialState : table -> state
exception Goto of state * nonterm
val mkLrTable : {actions : ((term,action) pairlist * action) array,
gotos : (nonterm,state) pairlist array,
numStates : int, numRules : int,
initialState : state} -> table
end
(* TOKEN: signature revealing the internal structure of a token. This signature
TOKEN distinct from the signature {parser name}_TOKENS produced by ML-Yacc.
The {parser name}_TOKENS structures contain some types and functions to
construct tokens from values and positions.
The representation of token was very carefully chosen here to allow the
polymorphic parser to work without knowing the types of semantic values
or line numbers.
This has had an impact on the TOKENS structure produced by SML-Yacc, which
is a structure parameter to lexer functors. We would like to have some
type 'a token which functions to construct tokens would create. A
constructor function for a integer token might be
INT: int * 'a * 'a -> 'a token.
This is not possible because we need to have tokens with the representation
given below for the polymorphic parser.
Thus our constructur functions for tokens have the form:
INT: int * 'a * 'a -> (svalue,'a) token
This in turn has had an impact on the signature that lexers for SML-Yacc
must match and the types that a user must declare in the user declarations
section of lexers.
*)
signature TOKEN =
sig
structure LrTable : LR_TABLE
datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
val sameToken : ('a,'b) token * ('a,'b) token -> bool
end
(* LR_PARSER: signature for a polymorphic LR parser *)
signature LR_PARSER =
sig
structure Stream: STREAM
structure LrTable : LR_TABLE
structure Token : TOKEN
sharing LrTable = Token.LrTable
exception ParseError
val parse : {table : LrTable.table,
lexer : ('_b,'_c) Token.token Stream.stream,
arg: 'arg,
saction : int *
'_c *
(LrTable.state * ('_b * '_c * '_c)) list *
'arg ->
LrTable.nonterm *
('_b * '_c * '_c) *
((LrTable.state *('_b * '_c * '_c)) list),
void : '_b,
ec : { is_keyword : LrTable.term -> bool,
noShift : LrTable.term -> bool,
preferred_change : (LrTable.term list * LrTable.term list) list,
errtermvalue : LrTable.term -> '_b,
showTerminal : LrTable.term -> string,
terms: LrTable.term list,
error : string * '_c * '_c -> unit
},
lookahead : int (* max amount of lookahead used in *)
(* error correction *)
} -> '_b *
(('_b,'_c) Token.token Stream.stream)
end
(* LEXER: a signature that most lexers produced for use with SML-Yacc's
output will match. The user is responsible for declaring type token,
type pos, and type svalue in the UserDeclarations section of a lexer.
Note that type token is abstract in the lexer. This allows SML-Yacc to
create a TOKENS signature for use with lexers produced by ML-Lex that
treats the type token abstractly. Lexers that are functors parametrized by
a Tokens structure matching a TOKENS signature cannot examine the structure
of tokens.
*)
signature LEXER =
sig
structure UserDeclarations :
sig
type ('a,'b) token
type pos
type svalue
end
val makeLexer : (int -> string) -> unit ->
(UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
end
(* ARG_LEXER: the %arg option of ML-Lex allows users to produce lexers which
also take an argument before yielding a function from unit to a token
*)
signature ARG_LEXER =
sig
structure UserDeclarations :
sig
type ('a,'b) token
type pos
type svalue
type arg
end
val makeLexer : (int -> string) -> UserDeclarations.arg -> unit ->
(UserDeclarations.svalue,UserDeclarations.pos) UserDeclarations.token
end
(* PARSER_DATA: the signature of ParserData structures in {parser name}LrValsFun
produced by SML-Yacc. All such structures match this signature.
The {parser name}LrValsFun produces a structure which contains all the values
except for the lexer needed to call the polymorphic parser mentioned
before.
*)
signature PARSER_DATA =
sig
(* the type of line numbers *)
type pos
(* the type of semantic values *)
type svalue
(* the type of the user-supplied argument to the parser *)
type arg
(* the intended type of the result of the parser. This value is
produced by applying extract from the structure Actions to the
final semantic value resultiing from a parse.
*)
type result
structure LrTable : LR_TABLE
structure Token : TOKEN
sharing Token.LrTable = LrTable
(* structure Actions contains the functions which mantain the
semantic values stack in the parser. Void is used to provide
a default value for the semantic stack.
*)
structure Actions :
sig
val actions : int * pos *
(LrTable.state * (svalue * pos * pos)) list * arg->
LrTable.nonterm * (svalue * pos * pos) *
((LrTable.state *(svalue * pos * pos)) list)
val void : svalue
val extract : svalue -> result
end
(* structure EC contains information used to improve error
recovery in an error-correcting parser *)
structure EC :
sig
val is_keyword : LrTable.term -> bool
val noShift : LrTable.term -> bool
val preferred_change : (LrTable.term list * LrTable.term list) list
val errtermvalue : LrTable.term -> svalue
val showTerminal : LrTable.term -> string
val terms: LrTable.term list
end
(* table is the LR table for the parser *)
val table : LrTable.table
end
(* signature PARSER is the signature that most user parsers created by
SML-Yacc will match.
*)
signature PARSER =
sig
structure Token : TOKEN
structure Stream : STREAM
exception ParseError
(* type pos is the type of line numbers *)
type pos
(* type result is the type of the result from the parser *)
type result
(* the type of the user-supplied argument to the parser *)
type arg
(* type svalue is the type of semantic values for the semantic value
stack
*)
type svalue
(* val makeLexer is used to create a stream of tokens for the parser *)
val makeLexer : (int -> string) ->
(svalue,pos) Token.token Stream.stream
(* val parse takes a stream of tokens and a function to print
errors and returns a value of type result and a stream containing
the unused tokens
*)
val parse : int * ((svalue,pos) Token.token Stream.stream) *
(string * pos * pos -> unit) * arg ->
result * (svalue,pos) Token.token Stream.stream
val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
bool
end
(* signature ARG_PARSER is the signature that will be matched by parsers whose
lexer takes an additional argument.
*)
signature ARG_PARSER =
sig
structure Token : TOKEN
structure Stream : STREAM
exception ParseError
type arg
type lexarg
type pos
type result
type svalue
val makeLexer : (int -> string) -> lexarg ->
(svalue,pos) Token.token Stream.stream
val parse : int * ((svalue,pos) Token.token Stream.stream) *
(string * pos * pos -> unit) * arg ->
result * (svalue,pos) Token.token Stream.stream
val sameToken : (svalue,pos) Token.token * (svalue,pos) Token.token ->
bool
end

View File

@ -0,0 +1,40 @@
This package was debianized by Aaron Matthew Read <amread@nyx.net> on
Fri, 25 Oct 2002 16:54:10 -0800.
It was downloaded from http://smlnj.cs.uchicago.edu/dist/working
Upstream Authors: The SML/NJ Team <smlnj-dev-list@lists.sourceforge.net>
Copyright: 2003-2008 The SML/NJ Fellowship
1989-2002 Lucent Technologies
1991-2003 John Reppy
1996-1998,2000 YALE FLINT PROJECT
1992 Vrije Universiteit, The Netherlands
1989-1992 Andrew W. Appel, James S. Mattson, David R. Tarditi
1988 Evans & Sutherland Computer Corporation, Salt Lake City, Utah
STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
Copyright (c) 1989-2002 by Lucent Technologies
Permission to use, copy, modify, and distribute this software and its
documentation for any purpose and without fee is hereby granted,
provided that the above copyright notice appear in all copies and that
both the copyright notice and this permission notice and warranty
disclaimer appear in supporting documentation, and that the name of
Lucent Technologies, Bell Labs or any Lucent entity not be used in
advertising or publicity pertaining to distribution of the software
without specific, written prior permission.
Lucent disclaims all warranties with regard to this software,
including all implied warranties of merchantability and fitness. In no
event shall Lucent be liable for any special, indirect or
consequential damages or any damages whatsoever resulting from loss of
use, data or profits, whether in an action of contract, negligence or
other tortious action, arising out of or in connection with the use
or performance of this software.
The SML/NJ distribution also includes code licensed under the same
terms as above, but with "David R. Tarditi Jr. and Andrew W. Appel",
"Vrije Universiteit" or "Evans & Sutherland" instead of "Lucent".

View File

@ -0,0 +1,118 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
(* functor Join creates a user parser by putting together a Lexer structure,
an LrValues structure, and a polymorphic parser structure. Note that
the Lexer and LrValues structure must share the type pos (i.e. the type
of line numbers), the type svalues for semantic values, and the type
of tokens.
*)
functor Join(structure Lex : LEXER
structure ParserData: PARSER_DATA
structure LrParser : LR_PARSER
sharing ParserData.LrTable = LrParser.LrTable
sharing ParserData.Token = LrParser.Token
sharing type Lex.UserDeclarations.svalue = ParserData.svalue
sharing type Lex.UserDeclarations.pos = ParserData.pos
sharing type Lex.UserDeclarations.token = ParserData.Token.token)
: PARSER =
struct
structure Token = ParserData.Token
structure Stream = LrParser.Stream
exception ParseError = LrParser.ParseError
type arg = ParserData.arg
type pos = ParserData.pos
type result = ParserData.result
type svalue = ParserData.svalue
val makeLexer = LrParser.Stream.streamify o Lex.makeLexer
val parse = fn (lookahead,lexer,error,arg) =>
(fn (a,b) => (ParserData.Actions.extract a,b))
(LrParser.parse {table = ParserData.table,
lexer=lexer,
lookahead=lookahead,
saction = ParserData.Actions.actions,
arg=arg,
void= ParserData.Actions.void,
ec = {is_keyword = ParserData.EC.is_keyword,
noShift = ParserData.EC.noShift,
preferred_change = ParserData.EC.preferred_change,
errtermvalue = ParserData.EC.errtermvalue,
error=error,
showTerminal = ParserData.EC.showTerminal,
terms = ParserData.EC.terms}}
)
val sameToken = Token.sameToken
end
(* functor JoinWithArg creates a variant of the parser structure produced
above. In this case, the makeLexer take an additional argument before
yielding a value of type unit -> (svalue,pos) token
*)
functor JoinWithArg(structure Lex : ARG_LEXER
structure ParserData: PARSER_DATA
structure LrParser : LR_PARSER
sharing ParserData.LrTable = LrParser.LrTable
sharing ParserData.Token = LrParser.Token
sharing type Lex.UserDeclarations.svalue = ParserData.svalue
sharing type Lex.UserDeclarations.pos = ParserData.pos
sharing type Lex.UserDeclarations.token = ParserData.Token.token)
: ARG_PARSER =
struct
structure Token = ParserData.Token
structure Stream = LrParser.Stream
exception ParseError = LrParser.ParseError
type arg = ParserData.arg
type lexarg = Lex.UserDeclarations.arg
type pos = ParserData.pos
type result = ParserData.result
type svalue = ParserData.svalue
val makeLexer = fn s => fn arg =>
LrParser.Stream.streamify (Lex.makeLexer s arg)
val parse = fn (lookahead,lexer,error,arg) =>
(fn (a,b) => (ParserData.Actions.extract a,b))
(LrParser.parse {table = ParserData.table,
lexer=lexer,
lookahead=lookahead,
saction = ParserData.Actions.actions,
arg=arg,
void= ParserData.Actions.void,
ec = {is_keyword = ParserData.EC.is_keyword,
noShift = ParserData.EC.noShift,
preferred_change = ParserData.EC.preferred_change,
errtermvalue = ParserData.EC.errtermvalue,
error=error,
showTerminal = ParserData.EC.showTerminal,
terms = ParserData.EC.terms}}
)
val sameToken = Token.sameToken
end;

View File

@ -0,0 +1,83 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
structure LrTable : LR_TABLE =
struct
open Array List
infix 9 sub
datatype ('a,'b) pairlist = EMPTY
| PAIR of 'a * 'b * ('a,'b) pairlist
datatype term = T of int
datatype nonterm = NT of int
datatype state = STATE of int
datatype action = SHIFT of state
| REDUCE of int (* rulenum from grammar *)
| ACCEPT
| ERROR
exception Goto of state * nonterm
type table = {states: int, rules : int,initialState: state,
action: ((term,action) pairlist * action) array,
goto : (nonterm,state) pairlist array}
val numStates = fn ({states,...} : table) => states
val numRules = fn ({rules,...} : table) => rules
val describeActions =
fn ({action,...} : table) =>
fn (STATE s) => action sub s
val describeGoto =
fn ({goto,...} : table) =>
fn (STATE s) => goto sub s
fun findTerm (T term,row,default) =
let fun find (PAIR (T key,data,r)) =
if key < term then find r
else if key=term then data
else default
| find EMPTY = default
in find row
end
fun findNonterm (NT nt,row) =
let fun find (PAIR (NT key,data,r)) =
if key < nt then find r
else if key=nt then SOME data
else NONE
| find EMPTY = NONE
in find row
end
val action = fn ({action,...} : table) =>
fn (STATE state,term) =>
let val (row,default) = action sub state
in findTerm(term,row,default)
end
val goto = fn ({goto,...} : table) =>
fn (a as (STATE state,nonterm)) =>
case findNonterm(nonterm,goto sub state)
of SOME state => state
| NONE => raise (Goto a)
val initialState = fn ({initialState,...} : table) => initialState
val mkLrTable = fn {actions,gotos,initialState,numStates,numRules} =>
({action=actions,goto=gotos,
states=numStates,
rules=numRules,
initialState=initialState} : table)
end;

View File

@ -0,0 +1,567 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
(* parser.sml: This is a parser driver for LR tables with an error-recovery
routine added to it. The routine used is described in detail in this
article:
'A Practical Method for LR and LL Syntactic Error Diagnosis and
Recovery', by M. Burke and G. Fisher, ACM Transactions on
Programming Langauges and Systems, Vol. 9, No. 2, April 1987,
pp. 164-197.
This program is an implementation is the partial, deferred method discussed
in the article. The algorithm and data structures used in the program
are described below.
This program assumes that all semantic actions are delayed. A semantic
action should produce a function from unit -> value instead of producing the
normal value. The parser returns the semantic value on the top of the
stack when accept is encountered. The user can deconstruct this value
and apply the unit -> value function in it to get the answer.
It also assumes that the lexer is a lazy stream.
Data Structures:
----------------
* The parser:
The state stack has the type
(state * (semantic value * line # * line #)) list
The parser keeps a queue of (state stack * lexer pair). A lexer pair
consists of a terminal * value pair and a lexer. This allows the
parser to reconstruct the states for terminals to the left of a
syntax error, and attempt to make error corrections there.
The queue consists of a pair of lists (x,y). New additions to
the queue are cons'ed onto y. The first element of x is the top
of the queue. If x is nil, then y is reversed and used
in place of x.
Algorithm:
----------
* The steady-state parser:
This parser keeps the length of the queue of state stacks at
a steady state by always removing an element from the front when
another element is placed on the end.
It has these arguments:
stack: current stack
queue: value of the queue
lexPair ((terminal,value),lex stream)
When SHIFT is encountered, the state to shift to and the value are
are pushed onto the state stack. The state stack and lexPair are
placed on the queue. The front element of the queue is removed.
When REDUCTION is encountered, the rule is applied to the current
stack to yield a triple (nonterm,value,new stack). A new
stack is formed by adding (goto(top state of stack,nonterm),value)
to the stack.
When ACCEPT is encountered, the top value from the stack and the
lexer are returned.
When an ERROR is encountered, fixError is called. FixError
takes the arguments to the parser, fixes the error if possible and
returns a new set of arguments.
* The distance-parser:
This parser includes an additional argument distance. It pushes
elements on the queue until it has parsed distance tokens, or an
ACCEPT or ERROR occurs. It returns a stack, lexer, the number of
tokens left unparsed, a queue, and an action option.
*)
signature FIFO =
sig type 'a queue
val empty : 'a queue
exception Empty
val get : 'a queue -> 'a * 'a queue
val put : 'a * 'a queue -> 'a queue
end
(* drt (12/15/89) -- the functor should be used in development work, but
it wastes space in the release version.
functor ParserGen(structure LrTable : LR_TABLE
structure Stream : STREAM) : LR_PARSER =
*)
structure LrParser :> LR_PARSER =
struct
structure LrTable = LrTable
structure Stream = Stream
val print = warning (* fn s => TextIO.output(TextIO.stdOut,s) *)
fun eqT (LrTable.T i, LrTable.T i') = i = i'
structure Token : TOKEN =
struct
structure LrTable = LrTable
datatype ('a,'b) token = TOKEN of LrTable.term * ('a * 'b * 'b)
val sameToken = fn (TOKEN(t,_),TOKEN(t',_)) => eqT (t,t')
end
open LrTable
open Token
val DEBUG1 = false
val DEBUG2 = false
exception ParseError
exception ParseImpossible of int
structure Fifo :> FIFO =
struct
type 'a queue = ('a list * 'a list)
val empty = (nil,nil)
exception Empty
fun get(a::x, y) = (a, (x,y))
| get(nil, nil) = raise Empty
| get(nil, y) = get(rev y, nil)
fun put(a,(x,y)) = (x,a::y)
end
type ('a,'b) elem = (state * ('a * 'b * 'b))
type ('a,'b) stack = ('a,'b) elem list
type ('a,'b) lexv = ('a,'b) token
type ('a,'b) lexpair = ('a,'b) lexv * (('a,'b) lexv Stream.stream)
type ('a,'b) distanceParse =
('a,'b) lexpair *
('a,'b) stack *
(('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
int ->
('a,'b) lexpair *
('a,'b) stack *
(('a,'b) stack * ('a,'b) lexpair) Fifo.queue *
int *
action option
type ('a,'b) ecRecord =
{is_keyword : term -> bool,
preferred_change : (term list * term list) list,
error : string * 'b * 'b -> unit,
errtermvalue : term -> 'a,
terms : term list,
showTerminal : term -> string,
noShift : term -> bool}
local
val print = warning (* fn s => TextIO.output(TextIO.stdOut,s) *)
val println = fn s => (print s; print "\n")
val showState = fn (STATE s) => "STATE " ^ (Int.toString s)
in
fun printStack(stack: ('a,'b) stack, n: int) =
case stack
of (state,_) :: rest =>
(print("\t" ^ Int.toString n ^ ": ");
println(showState state);
printStack(rest, n+1))
| nil => ()
fun prAction showTerminal
(stack as (state,_) :: _, next as (TOKEN (term,_),_), action) =
(println "Parse: state stack:";
printStack(stack, 0);
print(" state="
^ showState state
^ " next="
^ showTerminal term
^ " action="
);
case action
of SHIFT state => println ("SHIFT " ^ (showState state))
| REDUCE i => println ("REDUCE " ^ (Int.toString i))
| ERROR => println "ERROR"
| ACCEPT => println "ACCEPT")
| prAction _ (_,_,action) = ()
end
(* ssParse: parser which maintains the queue of (state * lexvalues) in a
steady-state. It takes a table, showTerminal function, saction
function, and fixError function. It parses until an ACCEPT is
encountered, or an exception is raised. When an error is encountered,
fixError is called with the arguments of parseStep (lexv,stack,and
queue). It returns the lexv, and a new stack and queue adjusted so
that the lexv can be parsed *)
val ssParse =
fn (table,showTerminal,saction,fixError,arg) =>
let val prAction = prAction showTerminal
val action = LrTable.action table
val goto = LrTable.goto table
fun parseStep(args as
(lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
lexer
),
stack as (state,_) :: _,
queue)) =
let val nextAction = action (state,terminal)
val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
else ()
in case nextAction
of SHIFT s =>
let val newStack = (s,value) :: stack
val newLexPair = Stream.get lexer
val (_,newQueue) =Fifo.get(Fifo.put((newStack,newLexPair),
queue))
in parseStep(newLexPair,(s,value)::stack,newQueue)
end
| REDUCE i =>
(case saction(i,leftPos,stack,arg)
of (nonterm,value,stack as (state,_) :: _) =>
parseStep(lexPair,(goto(state,nonterm),value)::stack,
queue)
| _ => raise (ParseImpossible 197))
| ERROR => parseStep(fixError args)
| ACCEPT =>
(case stack
of (_,(topvalue,_,_)) :: _ =>
let val (token,restLexer) = lexPair
in (topvalue,Stream.cons(token,restLexer))
end
| _ => raise (ParseImpossible 202))
end
| parseStep _ = raise (ParseImpossible 204)
in parseStep
end
(* distanceParse: parse until n tokens are shifted, or accept or
error are encountered. Takes a table, showTerminal function, and
semantic action function. Returns a parser which takes a lexPair
(lex result * lexer), a state stack, a queue, and a distance
(must be > 0) to parse. The parser returns a new lex-value, a stack
with the nth token shifted on top, a queue, a distance, and action
option. *)
val distanceParse =
fn (table,showTerminal,saction,arg) =>
let val prAction = prAction showTerminal
val action = LrTable.action table
val goto = LrTable.goto table
fun parseStep(lexPair,stack,queue,0) = (lexPair,stack,queue,0,NONE)
| parseStep(lexPair as (TOKEN (terminal, value as (_,leftPos,_)),
lexer
),
stack as (state,_) :: _,
queue,distance) =
let val nextAction = action(state,terminal)
val _ = if DEBUG1 then prAction(stack,lexPair,nextAction)
else ()
in case nextAction
of SHIFT s =>
let val newStack = (s,value) :: stack
val newLexPair = Stream.get lexer
in parseStep(newLexPair,(s,value)::stack,
Fifo.put((newStack,newLexPair),queue),distance-1)
end
| REDUCE i =>
(case saction(i,leftPos,stack,arg)
of (nonterm,value,stack as (state,_) :: _) =>
parseStep(lexPair,(goto(state,nonterm),value)::stack,
queue,distance)
| _ => raise (ParseImpossible 240))
| ERROR => (lexPair,stack,queue,distance,SOME nextAction)
| ACCEPT => (lexPair,stack,queue,distance,SOME nextAction)
end
| parseStep _ = raise (ParseImpossible 242)
in parseStep : ('_a,'_b) distanceParse
end
(* mkFixError: function to create fixError function which adjusts parser state
so that parse may continue in the presence of an error *)
fun mkFixError({is_keyword,terms,errtermvalue,
preferred_change,noShift,
showTerminal,error,...} : ('_a,'_b) ecRecord,
distanceParse : ('_a,'_b) distanceParse,
minAdvance,maxAdvance)
(lexv as (TOKEN (term,value as (_,leftPos,_)),_),stack,queue) =
let val _ = if DEBUG2 then
error("syntax error found at " ^ (showTerminal term),
leftPos,leftPos)
else ()
fun tokAt(t,p) = TOKEN(t,(errtermvalue t,p,p))
val minDelta = 3
(* pull all the state * lexv elements from the queue *)
val stateList =
let fun f q = let val (elem,newQueue) = Fifo.get q
in elem :: (f newQueue)
end handle Fifo.Empty => nil
in f queue
end
(* now number elements of stateList, giving distance from
error token *)
val (_, numStateList) =
List.foldr (fn (a,(num,r)) => (num+1,(a,num)::r)) (0, []) stateList
(* Represent the set of potential changes as a linked list.
Values of datatype Change hold information about a potential change.
oper = oper to be applied
pos = the # of the element in stateList that would be altered.
distance = the number of tokens beyond the error token which the
change allows us to parse.
new = new terminal * value pair at that point
orig = original terminal * value pair at the point being changed.
*)
datatype ('a,'b) change = CHANGE of
{pos : int, distance : int, leftPos: 'b, rightPos: 'b,
new : ('a,'b) lexv list, orig : ('a,'b) lexv list}
val showTerms = String.concat o map (fn TOKEN(t,_) => " " ^ showTerminal t)
val printChange = fn c =>
let val CHANGE {distance,new,orig,pos,...} = c
in (print ("{distance= " ^ (Int.toString distance));
print (",orig ="); print(showTerms orig);
print (",new ="); print(showTerms new);
print (",pos= " ^ (Int.toString pos));
print "}\n")
end
val printChangeList = app printChange
(* parse: given a lexPair, a stack, and the distance from the error
token, return the distance past the error token that we are able to parse.*)
fun parse (lexPair,stack,queuePos : int) =
case distanceParse(lexPair,stack,Fifo.empty,queuePos+maxAdvance+1)
of (_,_,_,distance,SOME ACCEPT) =>
if maxAdvance-distance-1 >= 0
then maxAdvance
else maxAdvance-distance-1
| (_,_,_,distance,_) => maxAdvance - distance - 1
(* catList: String.concatenate results of scanning list *)
fun catList l f = List.foldr (fn(a,r)=> f a @ r) [] l
fun keywordsDelta new = if List.exists (fn(TOKEN(t,_))=>is_keyword t) new
then minDelta else 0
fun tryChange{lex,stack,pos,leftPos,rightPos,orig,new} =
let val lex' = List.foldr (fn (t',p)=>(t',Stream.cons p)) lex new
val distance = parse(lex',stack,pos+length new-length orig)
in if distance >= minAdvance + keywordsDelta new
then [CHANGE{pos=pos,leftPos=leftPos,rightPos=rightPos,
distance=distance,orig=orig,new=new}]
else []
end
(* tryDelete: Try to delete n terminals.
Return single-element [success] or nil.
Do not delete unshiftable terminals. *)
fun tryDelete n ((stack,lexPair as (TOKEN(term,(_,l,r)),_)),qPos) =
let fun del(0,accum,left,right,lexPair) =
tryChange{lex=lexPair,stack=stack,
pos=qPos,leftPos=left,rightPos=right,
orig=rev accum, new=[]}
| del(n,accum,left,right,(tok as TOKEN(term,(_,_,r)),lexer)) =
if noShift term then []
else del(n-1,tok::accum,left,r,Stream.get lexer)
in del(n,[],l,r,lexPair)
end
(* tryInsert: try to insert tokens before the current terminal;
return a list of the successes *)
fun tryInsert((stack,lexPair as (TOKEN(_,(_,l,_)),_)),queuePos) =
catList terms (fn t =>
tryChange{lex=lexPair,stack=stack,
pos=queuePos,orig=[],new=[tokAt(t,l)],
leftPos=l,rightPos=l})
(* trySubst: try to substitute tokens for the current terminal;
return a list of the successes *)
fun trySubst ((stack,lexPair as (orig as TOKEN (term,(_,l,r)),lexer)),
queuePos) =
if noShift term then []
else
catList terms (fn t =>
tryChange{lex=Stream.get lexer,stack=stack,
pos=queuePos,
leftPos=l,rightPos=r,orig=[orig],
new=[tokAt(t,r)]})
(* do_delete(toks,lexPair) tries to delete tokens "toks" from "lexPair".
If it succeeds, returns SOME(toks',l,r,lp), where
toks' is the actual tokens (with positions and values) deleted,
(l,r) are the (leftmost,rightmost) position of toks',
lp is what remains of the stream after deletion
*)
fun do_delete(nil,lp as (TOKEN(_,(_,l,_)),_)) = SOME(nil,l,l,lp)
| do_delete([t],(tok as TOKEN(t',(_,l,r)),lp')) =
if eqT (t, t')
then SOME([tok],l,r,Stream.get lp')
else NONE
| do_delete(t::rest,(tok as TOKEN(t',(_,l,r)),lp')) =
if eqT (t,t')
then case do_delete(rest,Stream.get lp')
of SOME(deleted,l',r',lp'') =>
SOME(tok::deleted,l,r',lp'')
| NONE => NONE
else NONE
fun tryPreferred((stack,lexPair),queuePos) =
catList preferred_change (fn (delete,insert) =>
if List.exists noShift delete then [] (* should give warning at
parser-generation time *)
else case do_delete(delete,lexPair)
of SOME(deleted,l,r,lp) =>
tryChange{lex=lp,stack=stack,pos=queuePos,
leftPos=l,rightPos=r,orig=deleted,
new=map (fn t=>(tokAt(t,r))) insert}
| NONE => [])
val changes = catList numStateList tryPreferred @
catList numStateList tryInsert @
catList numStateList trySubst @
catList numStateList (tryDelete 1) @
catList numStateList (tryDelete 2) @
catList numStateList (tryDelete 3)
val findMaxDist = fn l =>
List.foldr (fn (CHANGE {distance,...},high) => Int.max(distance,high)) 0 l
(* maxDist: max distance past error taken that we could parse *)
val maxDist = findMaxDist changes
(* remove changes which did not parse maxDist tokens past the error token *)
val changes = catList changes
(fn(c as CHANGE{distance,...}) =>
if distance=maxDist then [c] else [])
in case changes
of (l as change :: _) =>
let fun print_msg (CHANGE {new,orig,leftPos,rightPos,...}) =
let val s =
case (orig,new)
of (_::_,[]) => "deleting " ^ (showTerms orig)
| ([],_::_) => "inserting " ^ (showTerms new)
| _ => "replacing " ^ (showTerms orig) ^
" with " ^ (showTerms new)
in error ("syntax error: " ^ s,leftPos,rightPos)
end
val _ =
(if length l > 1 andalso DEBUG2 then
(print "multiple fixes possible; could fix it by:\n";
app print_msg l;
print "chosen correction:\n")
else ();
print_msg change)
(* findNth: find nth queue entry from the error
entry. Returns the Nth queue entry and the portion of
the queue from the beginning to the nth-1 entry. The
error entry is at the end of the queue.
Examples:
queue = a b c d e
findNth 0 = (e,a b c d)
findNth 1 = (d,a b c)
*)
val findNth = fn n =>
let fun f (h::t,0) = (h,rev t)
| f (h::t,n) = f(t,n-1)
| f (nil,_) = let exception FindNth
in raise FindNth
end
in f (rev stateList,n)
end
val CHANGE {pos,orig,new,...} = change
val (last,queueFront) = findNth pos
val (stack,lexPair) = last
val lp1 = List.foldl(fn (_,(_,r)) => Stream.get r) lexPair orig
val lp2 = List.foldr(fn(t,r)=>(t,Stream.cons r)) lp1 new
val restQueue =
Fifo.put((stack,lp2),
List.foldl Fifo.put Fifo.empty queueFront)
val (lexPair,stack,queue,_,_) =
distanceParse(lp2,stack,restQueue,pos)
in (lexPair,stack,queue)
end
| nil => (error("syntax error found at " ^ (showTerminal term),
leftPos,leftPos); raise ParseError)
end
val parse = fn {arg,table,lexer,saction,void,lookahead,
ec=ec as {showTerminal,...} : ('_a,'_b) ecRecord} =>
let val distance = 15 (* defer distance tokens *)
val minAdvance = 1 (* must parse at least 1 token past error *)
val maxAdvance = Int.max(lookahead,0)(* max distance for parse check *)
val lexPair = Stream.get lexer
val (TOKEN (_,(_,leftPos,_)),_) = lexPair
val startStack = [(initialState table,(void,leftPos,leftPos))]
val startQueue = Fifo.put((startStack,lexPair),Fifo.empty)
val distanceParse = distanceParse(table,showTerminal,saction,arg)
val fixError = mkFixError(ec,distanceParse,minAdvance,maxAdvance)
val ssParse = ssParse(table,showTerminal,saction,fixError,arg)
fun loop (lexPair,stack,queue,_,SOME ACCEPT) =
ssParse(lexPair,stack,queue)
| loop (lexPair,stack,queue,0,_) = ssParse(lexPair,stack,queue)
| loop (lexPair,stack,queue,distance,SOME ERROR) =
let val (lexPair,stack,queue) = fixError(lexPair,stack,queue)
in loop (distanceParse(lexPair,stack,queue,distance))
end
| loop _ = let exception ParseInternal
in raise ParseInternal
end
in loop (distanceParse(lexPair,startStack,startQueue,distance))
end
end;

View File

@ -0,0 +1,29 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
use "base.sig";
use "join.sml";
use "lrtable.sml";
use "stream.sml";
use "parser2.sml";

View File

@ -0,0 +1,43 @@
(******************************************************************************
* STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
*
* Copyright (c) 1989-2002 by Lucent Technologies
*
* Permission to use, copy, modify, and distribute this software and its
* documentation for any purpose and without fee is hereby granted,
* provided that the above copyright notice appear in all copies and that
* both the copyright notice and this permission notice and warranty
* disclaimer appear in supporting documentation, and that the name of
* Lucent Technologies, Bell Labs or any Lucent entity not be used in
* advertising or publicity pertaining to distribution of the software
* without specific, written prior permission.
*
* Lucent disclaims all warranties with regard to this software,
* including all implied warranties of merchantability and fitness. In no
* event shall Lucent be liable for any special, indirect or
* consequential damages or any damages whatsoever resulting from loss of
* use, data or profits, whether in an action of contract, negligence or
* other tortious action, arising out of or in connection with the use
* or performance of this software.
******************************************************************************)
(* $Id$ *)
(* ML-Yacc Parser Generator (c) 1989 Andrew W. Appel, David R. Tarditi *)
(* Stream: a structure implementing a lazy stream. The signature STREAM
is found in base.sig *)
structure Stream :> STREAM =
struct
datatype 'a str = EVAL of 'a * 'a str Unsynchronized.ref | UNEVAL of (unit->'a)
type 'a stream = 'a str Unsynchronized.ref
fun get(Unsynchronized.ref(EVAL t)) = t
| get(s as Unsynchronized.ref(UNEVAL f)) =
let val t = (f(), Unsynchronized.ref(UNEVAL f)) in s := EVAL t; t end
fun streamify f = Unsynchronized.ref(UNEVAL f)
fun cons(a,s) = Unsynchronized.ref(EVAL(a,s))
end;

View File

@ -0,0 +1,101 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: ml_yacc_lib.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>ML Yacc Library\<close>
theory
"ml_yacc_lib"
imports
Main
begin
ML_file "ml-yacc-lib/base.sig"
ML_file "ml-yacc-lib/join.sml"
ML_file "ml-yacc-lib/lrtable.sml"
ML_file "ml-yacc-lib/stream.sml"
ML_file "ml-yacc-lib/parser2.sml"
(*
The files in the directory "ml-yacc-lib" are part of the sml/NJ language
processing tools. It was modified to work with Isabelle/ML by Achim D. Brucker.
It was downloaded from http://smlnj.cs.uchicago.edu/dist/working
Upstream Authors: The SML/NJ Team <smlnj-dev-list@lists.sourceforge.net>
Copyright: 2003-2008 The SML/NJ Fellowship
1989-2002 Lucent Technologies
1991-2003 John Reppy
1996-1998,2000 YALE FLINT PROJECT
1992 Vrije Universiteit, The Netherlands
1989-1992 Andrew W. Appel, James S. Mattson, David R. Tarditi
1988 Evans & Sutherland Computer Corporation, Salt Lake City, Utah
STANDARD ML OF NEW JERSEY COPYRIGHT NOTICE, LICENSE AND DISCLAIMER.
Copyright (c) 1989-2002 by Lucent Technologies
Permission to use, copy, modify, and distribute this software and its
documentation for any purpose and without fee is hereby granted,
provided that the above copyright notice appear in all copies and that
both the copyright notice and this permission notice and warranty
disclaimer appear in supporting documentation, and that the name of
Lucent Technologies, Bell Labs or any Lucent entity not be used in
advertising or publicity pertaining to distribution of the software
without specific, written prior permission.
Lucent disclaims all warranties with regard to this software,
including all implied warranties of merchantability and fitness. In no
event shall Lucent be liable for any special, indirect or
consequential damages or any damages whatsoever resulting from loss of
use, data or profits, whether in an action of contract, negligence or
other tortious action, arising out of or in connection with the use
or performance of this software.
The SML/NJ distribution also includes code licensed under the same
terms as above, but with "David R. Tarditi Jr. and Andrew W. Appel",
"Vrije Universiteit" or "Evans & Sutherland" instead of "Lucent".
*)
end

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,127 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: trac_fp_parser.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section\<open>Parser for Trac FP definitions\<close>
theory
trac_fp_parser
imports
"trac_term"
begin
ML_file "trac_parser/trac_fp.grm.sig"
ML_file "trac_parser/trac_fp.lex.sml"
ML_file "trac_parser/trac_fp.grm.sml"
ML\<open>
structure TracFpParser : sig
val parse_file: string -> (Trac_Term.cMsg) list
val parse_str: string -> (Trac_Term.cMsg) list
(* val term_of_trac: Trac_Term.cMsg -> term *)
val attack: Trac_Term.cMsg list -> bool
end =
struct
open Trac_Term
structure TracLrVals =
TracLrValsFun(structure Token = LrParser.Token)
structure TracLex =
TracLexFun(structure Tokens = TracLrVals.Tokens)
structure TracParser =
Join(structure LrParser = LrParser
structure ParserData = TracLrVals.ParserData
structure Lex = TracLex)
fun invoke lexstream =
let fun print_error (s,i:(int * int * int),_) =
TextIO.output(TextIO.stdOut,
"Error, line .... " ^ (Int.toString (#1 i)) ^"."^(Int.toString (#2 i ))^ ", " ^ s ^ "\n")
in TracParser.parse(0,lexstream,print_error,())
end
fun parse_fp lexer = let
val dummyEOF = TracLrVals.Tokens.EOF((0,0,0),(0,0,0))
fun certify (m,t) = Trac_Term.certifyMsg t [] m
fun loop lexer =
let
val _ = (TracLex.UserDeclarations.pos := (0,0,0);())
val (res,lexer) = invoke lexer
val (nextToken,lexer) = TracParser.Stream.get lexer
in if TracParser.sameToken(nextToken,dummyEOF) then ((),res)
else loop lexer
end
in map certify (#2(loop lexer))
end
fun parse_file tracFile = let
val infile = TextIO.openIn tracFile
val lexer = TracParser.makeLexer (fn _ => case ((TextIO.inputLine) infile) of
SOME s => s
| NONE => "")
in
parse_fp lexer
end
fun parse_str trac_fp_str = let
val parsed = Unsynchronized.ref false
fun input_string _ = if !parsed then "" else (parsed := true ;trac_fp_str)
val lexer = TracParser.makeLexer input_string
in
parse_fp lexer
end
fun attack fp = List.exists (fn e => e = cAttack) fp
(* fun term_of_trac (Trac_Term.cVar (n,t)) = @{const "cVar"}$(HOLogic.mk_tuple[HOLogic.mk_string n,
HOLogic.mk_string t])
| term_of_trac (Trac_Term.cConst n) = @{const "cConst"}$HOLogic.mk_string n
| term_of_trac (Trac_Term.cFun (n,l)) = @{const "cFun"}
$(HOLogic.mk_tuple[HOLogic.mk_string n, HOLogic.mk_list @{typ "cMsg"}
(map term_of_trac l)]) *)
end
\<close>
end

View File

@ -0,0 +1,126 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
open Trac_Term
exception NotYetSupported of string
%%
%eop EOF
%left
%name Trac
%term EOF
| COMMA of string
| FIXEDPOINT of string
| WHERE of string
| COLON of string
| PAREN_OPEN of string
| PAREN_CLOSE of string
| ASTERISK of string
| DOUBLE_ASTERISK of string
| DOUBLE_RARROW of string
| STRING_LITERAL of string
| UPPER_STRING_LITERAL of string
| LOWER_STRING_LITERAL of string
| INTEGER_LITERAL of string
| ONE of string
| ZERO of string
| ATTACK of string
%nonterm START of (Msg * TypeDecl list) list
| trac_file of (Msg * TypeDecl list) list
| symfact_list_exp of (Msg * TypeDecl list) list
| symfact_exp of Msg * TypeDecl list
| rule_exp of Msg
| arg_list_exp of Msg list
| arg_exp of Msg
| type_list_exp of TypeDecl list
| type_exp of TypeDecl
| string_literal of string
| upper_literal of string
| lower_literal of string
| int_literal of string
%pos (int * int * int)
%noshift EOF
%%
START: trac_file (trac_file)
trac_file: FIXEDPOINT symfact_list_exp (symfact_list_exp)
| symfact_list_exp (symfact_list_exp)
symfact_list_exp: symfact_exp ([symfact_exp])
| symfact_exp symfact_list_exp ([symfact_exp]@symfact_list_exp)
symfact_exp: DOUBLE_RARROW ATTACK ((Attack,[]))
| rule_exp WHERE type_list_exp ((rule_exp,type_list_exp))
| DOUBLE_RARROW rule_exp WHERE type_list_exp ((rule_exp,type_list_exp))
| DOUBLE_ASTERISK DOUBLE_RARROW rule_exp WHERE type_list_exp ((rule_exp,type_list_exp))
| rule_exp ((rule_exp,[]))
| DOUBLE_RARROW rule_exp ((rule_exp,[]))
| DOUBLE_ASTERISK DOUBLE_RARROW rule_exp ((rule_exp,[]))
rule_exp: upper_literal (Var (upper_literal))
| lower_literal (Fun (lower_literal,[]))
| lower_literal PAREN_OPEN arg_list_exp PAREN_CLOSE (Fun (lower_literal,arg_list_exp))
arg_list_exp: arg_exp ([arg_exp])
| arg_exp COMMA arg_list_exp ([arg_exp]@arg_list_exp)
arg_exp: rule_exp (rule_exp)
| ASTERISK int_literal (Var (int_literal))
| int_literal (Const (int_literal))
type_list_exp: type_exp ([type_exp])
| type_exp type_list_exp ([type_exp]@type_list_exp)
type_exp: ASTERISK int_literal COLON string_literal ((int_literal,string_literal))
| upper_literal COLON string_literal ((upper_literal,string_literal))
upper_literal: UPPER_STRING_LITERAL (UPPER_STRING_LITERAL)
lower_literal: LOWER_STRING_LITERAL (LOWER_STRING_LITERAL)
string_literal: upper_literal (upper_literal)
| lower_literal (lower_literal)
int_literal: INTEGER_LITERAL (INTEGER_LITERAL)
| ZERO ("0")
| ONE ("1")

View File

@ -0,0 +1,29 @@
signature Trac_TOKENS =
sig
type ('a,'b) token
type svalue
val ATTACK: (string) * 'a * 'a -> (svalue,'a) token
val ZERO: (string) * 'a * 'a -> (svalue,'a) token
val ONE: (string) * 'a * 'a -> (svalue,'a) token
val INTEGER_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val LOWER_STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val UPPER_STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val DOUBLE_RARROW: (string) * 'a * 'a -> (svalue,'a) token
val DOUBLE_ASTERISK: (string) * 'a * 'a -> (svalue,'a) token
val ASTERISK: (string) * 'a * 'a -> (svalue,'a) token
val PAREN_CLOSE: (string) * 'a * 'a -> (svalue,'a) token
val PAREN_OPEN: (string) * 'a * 'a -> (svalue,'a) token
val COLON: (string) * 'a * 'a -> (svalue,'a) token
val WHERE: (string) * 'a * 'a -> (svalue,'a) token
val FIXEDPOINT: (string) * 'a * 'a -> (svalue,'a) token
val COMMA: (string) * 'a * 'a -> (svalue,'a) token
val EOF: 'a * 'a -> (svalue,'a) token
end
signature Trac_LRVALS=
sig
structure Tokens : Trac_TOKENS
structure ParserData:PARSER_DATA
sharing type ParserData.Token.token = Tokens.token
sharing type ParserData.svalue = Tokens.svalue
end

View File

@ -0,0 +1,678 @@
(***** GENERATED FILE -- DO NOT EDIT ****)
functor TracLrValsFun(structure Token : TOKEN)
: sig structure ParserData : PARSER_DATA
structure Tokens : Trac_TOKENS
end
=
struct
structure ParserData=
struct
structure Header =
struct
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
open Trac_Term
exception NotYetSupported of string
end
structure LrTable = Token.LrTable
structure Token = Token
local open LrTable in
val table=let val actionRows =
"\
\\001\000\001\000\000\000\000\000\
\\001\000\003\000\013\000\009\000\012\000\010\000\011\000\012\000\010\000\
\\013\000\009\000\000\000\
\\001\000\005\000\038\000\000\000\
\\001\000\005\000\047\000\000\000\
\\001\000\007\000\036\000\000\000\
\\001\000\008\000\028\000\012\000\010\000\013\000\009\000\014\000\027\000\
\\015\000\026\000\016\000\025\000\000\000\
\\001\000\008\000\032\000\012\000\010\000\000\000\
\\001\000\009\000\012\000\010\000\011\000\012\000\010\000\013\000\009\000\000\000\
\\001\000\010\000\019\000\000\000\
\\001\000\012\000\010\000\013\000\009\000\000\000\
\\001\000\012\000\010\000\013\000\009\000\017\000\018\000\000\000\
\\001\000\014\000\027\000\015\000\026\000\016\000\025\000\000\000\
\\051\000\000\000\
\\052\000\000\000\
\\053\000\000\000\
\\054\000\009\000\012\000\010\000\011\000\012\000\010\000\013\000\009\000\000\000\
\\055\000\000\000\
\\056\000\000\000\
\\057\000\000\000\
\\058\000\000\000\
\\059\000\000\000\
\\060\000\004\000\015\000\000\000\
\\061\000\004\000\033\000\000\000\
\\062\000\004\000\042\000\000\000\
\\063\000\000\000\
\\064\000\006\000\014\000\000\000\
\\065\000\000\000\
\\066\000\002\000\035\000\000\000\
\\067\000\000\000\
\\068\000\000\000\
\\069\000\000\000\
\\070\000\000\000\
\\071\000\008\000\032\000\012\000\010\000\000\000\
\\072\000\000\000\
\\073\000\000\000\
\\074\000\000\000\
\\075\000\000\000\
\\076\000\000\000\
\\077\000\000\000\
\\078\000\000\000\
\\079\000\000\000\
\\080\000\000\000\
\\081\000\000\000\
\"
val actionRowNumbers =
"\001\000\025\000\024\000\021\000\
\\015\000\014\000\012\000\037\000\
\\036\000\010\000\008\000\007\000\
\\005\000\006\000\016\000\022\000\
\\017\000\009\000\013\000\031\000\
\\027\000\004\000\029\000\041\000\
\\042\000\040\000\011\000\002\000\
\\032\000\018\000\011\000\006\000\
\\023\000\005\000\026\000\030\000\
\\009\000\033\000\003\000\019\000\
\\006\000\028\000\039\000\038\000\
\\035\000\009\000\020\000\034\000\
\\000\000"
val gotoT =
"\
\\001\000\048\000\002\000\006\000\003\000\005\000\004\000\004\000\
\\005\000\003\000\011\000\002\000\012\000\001\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\003\000\014\000\004\000\004\000\005\000\003\000\011\000\002\000\
\\012\000\001\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\005\000\015\000\011\000\002\000\012\000\001\000\000\000\
\\000\000\
\\003\000\018\000\004\000\004\000\005\000\003\000\011\000\002\000\
\\012\000\001\000\000\000\
\\005\000\022\000\006\000\021\000\007\000\020\000\011\000\002\000\
\\012\000\001\000\013\000\019\000\000\000\
\\008\000\029\000\009\000\028\000\011\000\027\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\005\000\032\000\011\000\002\000\012\000\001\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\013\000\035\000\000\000\
\\000\000\
\\008\000\037\000\009\000\028\000\011\000\027\000\000\000\
\\000\000\
\\013\000\038\000\000\000\
\\008\000\039\000\009\000\028\000\011\000\027\000\000\000\
\\000\000\
\\005\000\022\000\006\000\041\000\007\000\020\000\011\000\002\000\
\\012\000\001\000\013\000\019\000\000\000\
\\000\000\
\\000\000\
\\010\000\044\000\011\000\043\000\012\000\042\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\008\000\046\000\009\000\028\000\011\000\027\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\\000\000\
\\010\000\047\000\011\000\043\000\012\000\042\000\000\000\
\\000\000\
\\000\000\
\\000\000\
\"
val numstates = 49
val numrules = 31
val s = Unsynchronized.ref "" and index = Unsynchronized.ref 0
val string_to_int = fn () =>
let val i = !index
in index := i+2; Char.ord(String.sub(!s,i)) + Char.ord(String.sub(!s,i+1)) * 256
end
val string_to_list = fn s' =>
let val len = String.size s'
fun f () =
if !index < len then string_to_int() :: f()
else nil
in index := 0; s := s'; f ()
end
val string_to_pairlist = fn (conv_key,conv_entry) =>
let fun f () =
case string_to_int()
of 0 => EMPTY
| n => PAIR(conv_key (n-1),conv_entry (string_to_int()),f())
in f
end
val string_to_pairlist_default = fn (conv_key,conv_entry) =>
let val conv_row = string_to_pairlist(conv_key,conv_entry)
in fn () =>
let val default = conv_entry(string_to_int())
val row = conv_row()
in (row,default)
end
end
val string_to_table = fn (convert_row,s') =>
let val len = String.size s'
fun f ()=
if !index < len then convert_row() :: f()
else nil
in (s := s'; index := 0; f ())
end
local
val memo = Array.array(numstates+numrules,ERROR)
val _ =let fun g i=(Array.update(memo,i,REDUCE(i-numstates)); g(i+1))
fun f i =
if i=numstates then g i
else (Array.update(memo,i,SHIFT (STATE i)); f (i+1))
in f 0 handle General.Subscript => ()
end
in
val entry_to_action = fn 0 => ACCEPT | 1 => ERROR | j => Array.sub(memo,(j-2))
end
val gotoT=Array.fromList(string_to_table(string_to_pairlist(NT,STATE),gotoT))
val actionRows=string_to_table(string_to_pairlist_default(T,entry_to_action),actionRows)
val actionRowNumbers = string_to_list actionRowNumbers
val actionT = let val actionRowLookUp=
let val a=Array.fromList(actionRows) in fn i=>Array.sub(a,i) end
in Array.fromList(List.map actionRowLookUp actionRowNumbers)
end
in LrTable.mkLrTable {actions=actionT,gotos=gotoT,numRules=numrules,
numStates=numstates,initialState=STATE 0}
end
end
local open Header in
type pos = ( int * int * int )
type arg = unit
structure MlyValue =
struct
datatype svalue = VOID | ntVOID of unit -> unit
| ATTACK of unit -> (string) | ZERO of unit -> (string)
| ONE of unit -> (string) | INTEGER_LITERAL of unit -> (string)
| LOWER_STRING_LITERAL of unit -> (string)
| UPPER_STRING_LITERAL of unit -> (string)
| STRING_LITERAL of unit -> (string)
| DOUBLE_RARROW of unit -> (string)
| DOUBLE_ASTERISK of unit -> (string)
| ASTERISK of unit -> (string) | PAREN_CLOSE of unit -> (string)
| PAREN_OPEN of unit -> (string) | COLON of unit -> (string)
| WHERE of unit -> (string) | FIXEDPOINT of unit -> (string)
| COMMA of unit -> (string) | int_literal of unit -> (string)
| lower_literal of unit -> (string)
| upper_literal of unit -> (string)
| string_literal of unit -> (string)
| type_exp of unit -> (TypeDecl)
| type_list_exp of unit -> (TypeDecl list)
| arg_exp of unit -> (Msg) | arg_list_exp of unit -> (Msg list)
| rule_exp of unit -> (Msg)
| symfact_exp of unit -> (Msg*TypeDecl list)
| symfact_list_exp of unit -> ( ( Msg * TypeDecl list ) list)
| trac_file of unit -> ( ( Msg * TypeDecl list ) list)
| START of unit -> ( ( Msg * TypeDecl list ) list)
end
type svalue = MlyValue.svalue
type result = ( Msg * TypeDecl list ) list
end
structure EC=
struct
open LrTable
infix 5 $$
fun x $$ y = y::x
val is_keyword =
fn _ => false
val preferred_change : (term list * term list) list =
nil
val noShift =
fn (T 0) => true | _ => false
val showTerminal =
fn (T 0) => "EOF"
| (T 1) => "COMMA"
| (T 2) => "FIXEDPOINT"
| (T 3) => "WHERE"
| (T 4) => "COLON"
| (T 5) => "PAREN_OPEN"
| (T 6) => "PAREN_CLOSE"
| (T 7) => "ASTERISK"
| (T 8) => "DOUBLE_ASTERISK"
| (T 9) => "DOUBLE_RARROW"
| (T 10) => "STRING_LITERAL"
| (T 11) => "UPPER_STRING_LITERAL"
| (T 12) => "LOWER_STRING_LITERAL"
| (T 13) => "INTEGER_LITERAL"
| (T 14) => "ONE"
| (T 15) => "ZERO"
| (T 16) => "ATTACK"
| _ => "bogus-term"
local open Header in
val errtermvalue=
fn _ => MlyValue.VOID
end
val terms : term list = nil
$$ (T 0)end
structure Actions =
struct
exception mlyAction of int
local open Header in
val actions =
fn (i392,defaultPos,stack,
(()):arg) =>
case (i392,stack)
of ( 0, ( ( _, ( MlyValue.trac_file trac_file1, trac_file1left,
trac_file1right)) :: rest671)) => let val result = MlyValue.START (fn
_ => let val (trac_file as trac_file1) = trac_file1 ()
in (trac_file)
end)
in ( LrTable.NT 0, ( result, trac_file1left, trac_file1right),
rest671)
end
| ( 1, ( ( _, ( MlyValue.symfact_list_exp symfact_list_exp1, _,
symfact_list_exp1right)) :: ( _, ( MlyValue.FIXEDPOINT FIXEDPOINT1,
FIXEDPOINT1left, _)) :: rest671)) => let val result =
MlyValue.trac_file (fn _ => let val FIXEDPOINT1 = FIXEDPOINT1 ()
val (symfact_list_exp as symfact_list_exp1) = symfact_list_exp1 ()
in (symfact_list_exp)
end)
in ( LrTable.NT 1, ( result, FIXEDPOINT1left, symfact_list_exp1right)
, rest671)
end
| ( 2, ( ( _, ( MlyValue.symfact_list_exp symfact_list_exp1,
symfact_list_exp1left, symfact_list_exp1right)) :: rest671)) => let
val result = MlyValue.trac_file (fn _ => let val (symfact_list_exp
as symfact_list_exp1) = symfact_list_exp1 ()
in (symfact_list_exp)
end)
in ( LrTable.NT 1, ( result, symfact_list_exp1left,
symfact_list_exp1right), rest671)
end
| ( 3, ( ( _, ( MlyValue.symfact_exp symfact_exp1, symfact_exp1left,
symfact_exp1right)) :: rest671)) => let val result =
MlyValue.symfact_list_exp (fn _ => let val (symfact_exp as
symfact_exp1) = symfact_exp1 ()
in ([symfact_exp])
end)
in ( LrTable.NT 2, ( result, symfact_exp1left, symfact_exp1right),
rest671)
end
| ( 4, ( ( _, ( MlyValue.symfact_list_exp symfact_list_exp1, _,
symfact_list_exp1right)) :: ( _, ( MlyValue.symfact_exp symfact_exp1,
symfact_exp1left, _)) :: rest671)) => let val result =
MlyValue.symfact_list_exp (fn _ => let val (symfact_exp as
symfact_exp1) = symfact_exp1 ()
val (symfact_list_exp as symfact_list_exp1) = symfact_list_exp1 ()
in ([symfact_exp]@symfact_list_exp)
end)
in ( LrTable.NT 2, ( result, symfact_exp1left, symfact_list_exp1right
), rest671)
end
| ( 5, ( ( _, ( MlyValue.ATTACK ATTACK1, _, ATTACK1right)) :: ( _, (
MlyValue.DOUBLE_RARROW DOUBLE_RARROW1, DOUBLE_RARROW1left, _)) ::
rest671)) => let val result = MlyValue.symfact_exp (fn _ => let val
DOUBLE_RARROW1 = DOUBLE_RARROW1 ()
val ATTACK1 = ATTACK1 ()
in ((Attack,[]))
end)
in ( LrTable.NT 3, ( result, DOUBLE_RARROW1left, ATTACK1right),
rest671)
end
| ( 6, ( ( _, ( MlyValue.type_list_exp type_list_exp1, _,
type_list_exp1right)) :: ( _, ( MlyValue.WHERE WHERE1, _, _)) :: ( _,
( MlyValue.rule_exp rule_exp1, rule_exp1left, _)) :: rest671)) => let
val result = MlyValue.symfact_exp (fn _ => let val (rule_exp as
rule_exp1) = rule_exp1 ()
val WHERE1 = WHERE1 ()
val (type_list_exp as type_list_exp1) = type_list_exp1 ()
in ((rule_exp,type_list_exp))
end)
in ( LrTable.NT 3, ( result, rule_exp1left, type_list_exp1right),
rest671)
end
| ( 7, ( ( _, ( MlyValue.type_list_exp type_list_exp1, _,
type_list_exp1right)) :: ( _, ( MlyValue.WHERE WHERE1, _, _)) :: ( _,
( MlyValue.rule_exp rule_exp1, _, _)) :: ( _, ( MlyValue.DOUBLE_RARROW
DOUBLE_RARROW1, DOUBLE_RARROW1left, _)) :: rest671)) => let val
result = MlyValue.symfact_exp (fn _ => let val DOUBLE_RARROW1 =
DOUBLE_RARROW1 ()
val (rule_exp as rule_exp1) = rule_exp1 ()
val WHERE1 = WHERE1 ()
val (type_list_exp as type_list_exp1) = type_list_exp1 ()
in ((rule_exp,type_list_exp))
end)
in ( LrTable.NT 3, ( result, DOUBLE_RARROW1left, type_list_exp1right)
, rest671)
end
| ( 8, ( ( _, ( MlyValue.type_list_exp type_list_exp1, _,
type_list_exp1right)) :: ( _, ( MlyValue.WHERE WHERE1, _, _)) :: ( _,
( MlyValue.rule_exp rule_exp1, _, _)) :: ( _, ( MlyValue.DOUBLE_RARROW
DOUBLE_RARROW1, _, _)) :: ( _, ( MlyValue.DOUBLE_ASTERISK
DOUBLE_ASTERISK1, DOUBLE_ASTERISK1left, _)) :: rest671)) => let val
result = MlyValue.symfact_exp (fn _ => let val DOUBLE_ASTERISK1 =
DOUBLE_ASTERISK1 ()
val DOUBLE_RARROW1 = DOUBLE_RARROW1 ()
val (rule_exp as rule_exp1) = rule_exp1 ()
val WHERE1 = WHERE1 ()
val (type_list_exp as type_list_exp1) = type_list_exp1 ()
in ((rule_exp,type_list_exp))
end)
in ( LrTable.NT 3, ( result, DOUBLE_ASTERISK1left,
type_list_exp1right), rest671)
end
| ( 9, ( ( _, ( MlyValue.rule_exp rule_exp1, rule_exp1left,
rule_exp1right)) :: rest671)) => let val result =
MlyValue.symfact_exp (fn _ => let val (rule_exp as rule_exp1) =
rule_exp1 ()
in ((rule_exp,[]))
end)
in ( LrTable.NT 3, ( result, rule_exp1left, rule_exp1right), rest671)
end
| ( 10, ( ( _, ( MlyValue.rule_exp rule_exp1, _, rule_exp1right)) ::
( _, ( MlyValue.DOUBLE_RARROW DOUBLE_RARROW1, DOUBLE_RARROW1left, _))
:: rest671)) => let val result = MlyValue.symfact_exp (fn _ => let
val DOUBLE_RARROW1 = DOUBLE_RARROW1 ()
val (rule_exp as rule_exp1) = rule_exp1 ()
in ((rule_exp,[]))
end)
in ( LrTable.NT 3, ( result, DOUBLE_RARROW1left, rule_exp1right),
rest671)
end
| ( 11, ( ( _, ( MlyValue.rule_exp rule_exp1, _, rule_exp1right)) ::
( _, ( MlyValue.DOUBLE_RARROW DOUBLE_RARROW1, _, _)) :: ( _, (
MlyValue.DOUBLE_ASTERISK DOUBLE_ASTERISK1, DOUBLE_ASTERISK1left, _))
:: rest671)) => let val result = MlyValue.symfact_exp (fn _ => let
val DOUBLE_ASTERISK1 = DOUBLE_ASTERISK1 ()
val DOUBLE_RARROW1 = DOUBLE_RARROW1 ()
val (rule_exp as rule_exp1) = rule_exp1 ()
in ((rule_exp,[]))
end)
in ( LrTable.NT 3, ( result, DOUBLE_ASTERISK1left, rule_exp1right),
rest671)
end
| ( 12, ( ( _, ( MlyValue.upper_literal upper_literal1,
upper_literal1left, upper_literal1right)) :: rest671)) => let val
result = MlyValue.rule_exp (fn _ => let val (upper_literal as
upper_literal1) = upper_literal1 ()
in (Var (upper_literal))
end)
in ( LrTable.NT 4, ( result, upper_literal1left, upper_literal1right)
, rest671)
end
| ( 13, ( ( _, ( MlyValue.lower_literal lower_literal1,
lower_literal1left, lower_literal1right)) :: rest671)) => let val
result = MlyValue.rule_exp (fn _ => let val (lower_literal as
lower_literal1) = lower_literal1 ()
in (Fun (lower_literal,[]))
end)
in ( LrTable.NT 4, ( result, lower_literal1left, lower_literal1right)
, rest671)
end
| ( 14, ( ( _, ( MlyValue.PAREN_CLOSE PAREN_CLOSE1, _,
PAREN_CLOSE1right)) :: ( _, ( MlyValue.arg_list_exp arg_list_exp1, _,
_)) :: ( _, ( MlyValue.PAREN_OPEN PAREN_OPEN1, _, _)) :: ( _, (
MlyValue.lower_literal lower_literal1, lower_literal1left, _)) ::
rest671)) => let val result = MlyValue.rule_exp (fn _ => let val (
lower_literal as lower_literal1) = lower_literal1 ()
val PAREN_OPEN1 = PAREN_OPEN1 ()
val (arg_list_exp as arg_list_exp1) = arg_list_exp1 ()
val PAREN_CLOSE1 = PAREN_CLOSE1 ()
in (Fun (lower_literal,arg_list_exp))
end)
in ( LrTable.NT 4, ( result, lower_literal1left, PAREN_CLOSE1right),
rest671)
end
| ( 15, ( ( _, ( MlyValue.arg_exp arg_exp1, arg_exp1left,
arg_exp1right)) :: rest671)) => let val result =
MlyValue.arg_list_exp (fn _ => let val (arg_exp as arg_exp1) =
arg_exp1 ()
in ([arg_exp])
end)
in ( LrTable.NT 5, ( result, arg_exp1left, arg_exp1right), rest671)
end
| ( 16, ( ( _, ( MlyValue.arg_list_exp arg_list_exp1, _,
arg_list_exp1right)) :: ( _, ( MlyValue.COMMA COMMA1, _, _)) :: ( _, (
MlyValue.arg_exp arg_exp1, arg_exp1left, _)) :: rest671)) => let val
result = MlyValue.arg_list_exp (fn _ => let val (arg_exp as arg_exp1
) = arg_exp1 ()
val COMMA1 = COMMA1 ()
val (arg_list_exp as arg_list_exp1) = arg_list_exp1 ()
in ([arg_exp]@arg_list_exp)
end)
in ( LrTable.NT 5, ( result, arg_exp1left, arg_list_exp1right),
rest671)
end
| ( 17, ( ( _, ( MlyValue.rule_exp rule_exp1, rule_exp1left,
rule_exp1right)) :: rest671)) => let val result = MlyValue.arg_exp
(fn _ => let val (rule_exp as rule_exp1) = rule_exp1 ()
in (rule_exp)
end)
in ( LrTable.NT 6, ( result, rule_exp1left, rule_exp1right), rest671)
end
| ( 18, ( ( _, ( MlyValue.int_literal int_literal1, _,
int_literal1right)) :: ( _, ( MlyValue.ASTERISK ASTERISK1,
ASTERISK1left, _)) :: rest671)) => let val result = MlyValue.arg_exp
(fn _ => let val ASTERISK1 = ASTERISK1 ()
val (int_literal as int_literal1) = int_literal1 ()
in (Var (int_literal))
end)
in ( LrTable.NT 6, ( result, ASTERISK1left, int_literal1right),
rest671)
end
| ( 19, ( ( _, ( MlyValue.int_literal int_literal1, int_literal1left,
int_literal1right)) :: rest671)) => let val result =
MlyValue.arg_exp (fn _ => let val (int_literal as int_literal1) =
int_literal1 ()
in (Const (int_literal))
end)
in ( LrTable.NT 6, ( result, int_literal1left, int_literal1right),
rest671)
end
| ( 20, ( ( _, ( MlyValue.type_exp type_exp1, type_exp1left,
type_exp1right)) :: rest671)) => let val result =
MlyValue.type_list_exp (fn _ => let val (type_exp as type_exp1) =
type_exp1 ()
in ([type_exp])
end)
in ( LrTable.NT 7, ( result, type_exp1left, type_exp1right), rest671)
end
| ( 21, ( ( _, ( MlyValue.type_list_exp type_list_exp1, _,
type_list_exp1right)) :: ( _, ( MlyValue.type_exp type_exp1,
type_exp1left, _)) :: rest671)) => let val result =
MlyValue.type_list_exp (fn _ => let val (type_exp as type_exp1) =
type_exp1 ()
val (type_list_exp as type_list_exp1) = type_list_exp1 ()
in ([type_exp]@type_list_exp)
end)
in ( LrTable.NT 7, ( result, type_exp1left, type_list_exp1right),
rest671)
end
| ( 22, ( ( _, ( MlyValue.string_literal string_literal1, _,
string_literal1right)) :: ( _, ( MlyValue.COLON COLON1, _, _)) :: ( _,
( MlyValue.int_literal int_literal1, _, _)) :: ( _, (
MlyValue.ASTERISK ASTERISK1, ASTERISK1left, _)) :: rest671)) => let
val result = MlyValue.type_exp (fn _ => let val ASTERISK1 =
ASTERISK1 ()
val (int_literal as int_literal1) = int_literal1 ()
val COLON1 = COLON1 ()
val (string_literal as string_literal1) = string_literal1 ()
in ((int_literal,string_literal))
end)
in ( LrTable.NT 8, ( result, ASTERISK1left, string_literal1right),
rest671)
end
| ( 23, ( ( _, ( MlyValue.string_literal string_literal1, _,
string_literal1right)) :: ( _, ( MlyValue.COLON COLON1, _, _)) :: ( _,
( MlyValue.upper_literal upper_literal1, upper_literal1left, _)) ::
rest671)) => let val result = MlyValue.type_exp (fn _ => let val (
upper_literal as upper_literal1) = upper_literal1 ()
val COLON1 = COLON1 ()
val (string_literal as string_literal1) = string_literal1 ()
in ((upper_literal,string_literal))
end)
in ( LrTable.NT 8, ( result, upper_literal1left, string_literal1right
), rest671)
end
| ( 24, ( ( _, ( MlyValue.UPPER_STRING_LITERAL UPPER_STRING_LITERAL1,
UPPER_STRING_LITERAL1left, UPPER_STRING_LITERAL1right)) :: rest671))
=> let val result = MlyValue.upper_literal (fn _ => let val (
UPPER_STRING_LITERAL as UPPER_STRING_LITERAL1) = UPPER_STRING_LITERAL1
()
in (UPPER_STRING_LITERAL)
end)
in ( LrTable.NT 10, ( result, UPPER_STRING_LITERAL1left,
UPPER_STRING_LITERAL1right), rest671)
end
| ( 25, ( ( _, ( MlyValue.LOWER_STRING_LITERAL LOWER_STRING_LITERAL1,
LOWER_STRING_LITERAL1left, LOWER_STRING_LITERAL1right)) :: rest671))
=> let val result = MlyValue.lower_literal (fn _ => let val (
LOWER_STRING_LITERAL as LOWER_STRING_LITERAL1) = LOWER_STRING_LITERAL1
()
in (LOWER_STRING_LITERAL)
end)
in ( LrTable.NT 11, ( result, LOWER_STRING_LITERAL1left,
LOWER_STRING_LITERAL1right), rest671)
end
| ( 26, ( ( _, ( MlyValue.upper_literal upper_literal1,
upper_literal1left, upper_literal1right)) :: rest671)) => let val
result = MlyValue.string_literal (fn _ => let val (upper_literal as
upper_literal1) = upper_literal1 ()
in (upper_literal)
end)
in ( LrTable.NT 9, ( result, upper_literal1left, upper_literal1right)
, rest671)
end
| ( 27, ( ( _, ( MlyValue.lower_literal lower_literal1,
lower_literal1left, lower_literal1right)) :: rest671)) => let val
result = MlyValue.string_literal (fn _ => let val (lower_literal as
lower_literal1) = lower_literal1 ()
in (lower_literal)
end)
in ( LrTable.NT 9, ( result, lower_literal1left, lower_literal1right)
, rest671)
end
| ( 28, ( ( _, ( MlyValue.INTEGER_LITERAL INTEGER_LITERAL1,
INTEGER_LITERAL1left, INTEGER_LITERAL1right)) :: rest671)) => let val
result = MlyValue.int_literal (fn _ => let val (INTEGER_LITERAL as
INTEGER_LITERAL1) = INTEGER_LITERAL1 ()
in (INTEGER_LITERAL)
end)
in ( LrTable.NT 12, ( result, INTEGER_LITERAL1left,
INTEGER_LITERAL1right), rest671)
end
| ( 29, ( ( _, ( MlyValue.ZERO ZERO1, ZERO1left, ZERO1right)) ::
rest671)) => let val result = MlyValue.int_literal (fn _ => let val
ZERO1 = ZERO1 ()
in ("0")
end)
in ( LrTable.NT 12, ( result, ZERO1left, ZERO1right), rest671)
end
| ( 30, ( ( _, ( MlyValue.ONE ONE1, ONE1left, ONE1right)) :: rest671)
) => let val result = MlyValue.int_literal (fn _ => let val ONE1 =
ONE1 ()
in ("1")
end)
in ( LrTable.NT 12, ( result, ONE1left, ONE1right), rest671)
end
| _ => raise (mlyAction i392)
end
val void = MlyValue.VOID
val extract = fn a => (fn MlyValue.START x => x
| _ => let exception ParseInternal
in raise ParseInternal end) a ()
end
end
structure Tokens : Trac_TOKENS =
struct
type svalue = ParserData.svalue
type ('a,'b) token = ('a,'b) Token.token
fun EOF (p1,p2) = Token.TOKEN (ParserData.LrTable.T 0,(
ParserData.MlyValue.VOID,p1,p2))
fun COMMA (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 1,(
ParserData.MlyValue.COMMA (fn () => i),p1,p2))
fun FIXEDPOINT (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 2,(
ParserData.MlyValue.FIXEDPOINT (fn () => i),p1,p2))
fun WHERE (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 3,(
ParserData.MlyValue.WHERE (fn () => i),p1,p2))
fun COLON (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 4,(
ParserData.MlyValue.COLON (fn () => i),p1,p2))
fun PAREN_OPEN (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 5,(
ParserData.MlyValue.PAREN_OPEN (fn () => i),p1,p2))
fun PAREN_CLOSE (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 6,(
ParserData.MlyValue.PAREN_CLOSE (fn () => i),p1,p2))
fun ASTERISK (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 7,(
ParserData.MlyValue.ASTERISK (fn () => i),p1,p2))
fun DOUBLE_ASTERISK (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 8,(
ParserData.MlyValue.DOUBLE_ASTERISK (fn () => i),p1,p2))
fun DOUBLE_RARROW (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 9,(
ParserData.MlyValue.DOUBLE_RARROW (fn () => i),p1,p2))
fun STRING_LITERAL (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 10,(
ParserData.MlyValue.STRING_LITERAL (fn () => i),p1,p2))
fun UPPER_STRING_LITERAL (i,p1,p2) = Token.TOKEN (
ParserData.LrTable.T 11,(ParserData.MlyValue.UPPER_STRING_LITERAL
(fn () => i),p1,p2))
fun LOWER_STRING_LITERAL (i,p1,p2) = Token.TOKEN (
ParserData.LrTable.T 12,(ParserData.MlyValue.LOWER_STRING_LITERAL
(fn () => i),p1,p2))
fun INTEGER_LITERAL (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 13,(
ParserData.MlyValue.INTEGER_LITERAL (fn () => i),p1,p2))
fun ONE (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 14,(
ParserData.MlyValue.ONE (fn () => i),p1,p2))
fun ZERO (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 15,(
ParserData.MlyValue.ZERO (fn () => i),p1,p2))
fun ATTACK (i,p1,p2) = Token.TOKEN (ParserData.LrTable.T 16,(
ParserData.MlyValue.ATTACK (fn () => i),p1,p2))
end
end

View File

@ -0,0 +1,103 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
structure Tokens = Tokens
open Trac_Term
type pos = int * int * int
type svalue = Tokens.svalue
type ('a,'b) token = ('a,'b) Tokens.token
type lexresult= (svalue,pos) token
val pos = ref (0,0,0)
fun eof () = Tokens.EOF((!pos,!pos))
fun error (e,p : (int * int * int),_) = TextIO.output (TextIO.stdOut,
String.concat[
"line ", (Int.toString (#1 p)), "/",
(Int.toString (#2 p - #3 p)),": ", e, "\n"
])
fun inputPos yypos = ((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))),
(#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))))
fun inputPos_half yypos = (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))
%%
%header (functor TracLexFun(structure Tokens: Trac_TOKENS));
alpha=[A-Za-z_];
upper=[A-Z];
lower=[a-z];
digit=[0-9];
ws = [\ \t];
%%
\n => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex());
{ws}+ => (pos := (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))); lex());
(#)[^\n]*\n => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex());
"/*""/"*([^*/]|[^*]"/"|"*"[^/])*"*"*"*/" => (lex());
"," => (Tokens.COMMA(yytext,inputPos_half yypos,inputPos_half yypos));
"Fixedpoint" => (Tokens.FIXEDPOINT(yytext,inputPos_half yypos,inputPos_half yypos));
"where" => (Tokens.WHERE(yytext,inputPos_half yypos,inputPos_half yypos));
":" => (Tokens.COLON(yytext,inputPos_half yypos,inputPos_half yypos));
"(" => (Tokens.PAREN_OPEN(yytext,inputPos_half yypos,inputPos_half yypos));
")" => (Tokens.PAREN_CLOSE(yytext,inputPos_half yypos,inputPos_half yypos));
"**" => (Tokens.DOUBLE_ASTERISK(yytext,inputPos_half yypos,inputPos_half yypos));
"*" => (Tokens.ASTERISK(yytext,inputPos_half yypos,inputPos_half yypos));
"=>" => (Tokens.DOUBLE_RARROW(yytext,inputPos_half yypos,inputPos_half yypos));
"one" => (Tokens.ONE(yytext,inputPos_half yypos,inputPos_half yypos));
"zero" => (Tokens.ZERO(yytext,inputPos_half yypos,inputPos_half yypos));
"attack" => (Tokens.ATTACK(yytext,inputPos_half yypos,inputPos_half yypos));
{digit}+ => (Tokens.INTEGER_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
"'"({alpha}|{ws}|{digit})*(("."|"_"|"/"|"-")*({alpha}|{ws}|{digit})*)*"'" => (Tokens.STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
{upper}({alpha}|{digit})*("'")* => (Tokens.UPPER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
{lower}({alpha}|{digit})*("'")* => (Tokens.LOWER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
. => (error ("ignoring bad character "^yytext,
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))),
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))));
lex());

View File

@ -0,0 +1,728 @@
(***** GENERATED FILE -- DO NOT EDIT ****)
functor TracLexFun(structure Tokens: Trac_TOKENS)=
struct
structure UserDeclarations =
struct
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
structure Tokens = Tokens
open Trac_Term
type pos = int * int * int
type svalue = Tokens.svalue
type ('a,'b) token = ('a,'b) Tokens.token
type lexresult= (svalue,pos) token
val pos = Unsynchronized.ref (0,0,0)
fun eof () = Tokens.EOF((!pos,!pos))
fun error (e,p : (int * int * int),_) = TextIO.output (TextIO.stdOut,
String.concat[
"line ", (Int.toString (#1 p)), "/",
(Int.toString (#2 p - #3 p)),": ", e, "\n"
])
fun inputPos yypos = ((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))),
(#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))))
fun inputPos_half yypos = (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))
end (* end of user routines *)
exception LexError (* raised if illegal leaf action tried *)
structure Internal =
struct
datatype yyfinstate = N of int
type statedata = {fin : yyfinstate list, trans: string}
(* transition & final state table *)
val tab = let
val s = [
(0,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(1,
"\003\003\003\003\003\003\003\003\003\065\067\003\003\003\003\003\
\\003\003\003\003\003\003\003\003\003\003\003\003\003\003\003\003\
\\065\003\003\062\003\003\003\058\057\056\054\003\053\003\003\043\
\\041\041\041\041\041\041\041\041\041\041\040\003\003\038\003\003\
\\003\025\025\025\025\025\028\025\025\025\025\025\025\025\025\025\
\\025\025\025\025\025\025\025\025\025\025\025\003\003\003\003\003\
\\003\019\010\010\010\010\010\010\010\010\010\010\010\010\010\016\
\\010\010\010\010\010\010\010\011\010\010\004\003\003\003\003\003\
\\003"
),
(4,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\007\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(5,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(6,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(7,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\008\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(8,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\009\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(11,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\012\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(12,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\013\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(13,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\014\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(14,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\015\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(16,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\017\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(17,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\018\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(19,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\020\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(20,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\021\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(21,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\022\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(22,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\023\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(23,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\006\000\000\000\000\000\000\000\000\
\\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\000\
\\000\005\005\005\005\005\005\005\005\005\005\005\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\005\
\\000\005\005\005\005\005\005\005\005\005\005\024\005\005\005\005\
\\005\005\005\005\005\005\005\005\005\005\005\000\000\000\000\000\
\\000"
),
(25,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(27,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(28,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\029\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(29,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\030\026\026\000\000\000\000\000\
\\000"
),
(30,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\031\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(31,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\032\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(32,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\033\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(33,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\034\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(34,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\035\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(35,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\036\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(36,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\027\000\000\000\000\000\000\000\000\
\\026\026\026\026\026\026\026\026\026\026\000\000\000\000\000\000\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\026\026\026\026\026\026\026\000\000\000\000\026\
\\000\026\026\026\026\026\026\026\026\026\026\026\026\026\026\026\
\\026\026\026\026\037\026\026\026\026\026\026\000\000\000\000\000\
\\000"
),
(38,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\039\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(41,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\042\042\042\042\042\042\042\042\042\042\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(43,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\044\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(44,
"\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\046\045\045\045\045\052\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045"
),
(45,
"\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\046\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045"
),
(46,
"\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\050\047\047\047\047\049\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\047\
\\047"
),
(47,
"\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\046\045\045\045\045\048\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045"
),
(48,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\047\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(50,
"\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\046\045\045\045\045\051\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\045\
\\045"
),
(54,
"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\055\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(58,
"\000\000\000\000\000\000\000\000\000\059\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\059\000\000\000\000\000\000\061\000\000\000\000\000\060\060\060\
\\059\059\059\059\059\059\059\059\059\059\000\000\000\000\000\000\
\\000\059\059\059\059\059\059\059\059\059\059\059\059\059\059\059\
\\059\059\059\059\059\059\059\059\059\059\059\000\000\000\000\059\
\\000\059\059\059\059\059\059\059\059\059\059\059\059\059\059\059\
\\059\059\059\059\059\059\059\059\059\059\059\000\000\000\000\000\
\\000"
),
(60,
"\000\000\000\000\000\000\000\000\000\060\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\060\000\000\000\000\000\000\061\000\000\000\000\000\060\060\060\
\\060\060\060\060\060\060\060\060\060\060\000\000\000\000\000\000\
\\000\060\060\060\060\060\060\060\060\060\060\060\060\060\060\060\
\\060\060\060\060\060\060\060\060\060\060\060\000\000\000\000\060\
\\000\060\060\060\060\060\060\060\060\060\060\060\060\060\060\060\
\\060\060\060\060\060\060\060\060\060\060\060\000\000\000\000\000\
\\000"
),
(62,
"\063\063\063\063\063\063\063\063\063\063\064\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\063\
\\063"
),
(65,
"\000\000\000\000\000\000\000\000\000\066\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\066\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
\\000"
),
(0, "")]
fun f x = x
val s = List.map f (List.rev (tl (List.rev s)))
exception LexHackingError
fun look ((j,x)::r, i: int) = if i = j then x else look(r, i)
| look ([], i) = raise LexHackingError
fun g {fin=x, trans=i} = {fin=x, trans=look(s,i)}
in Vector.fromList(List.map g
[{fin = [], trans = 0},
{fin = [], trans = 1},
{fin = [], trans = 1},
{fin = [(N 97)], trans = 0},
{fin = [(N 95),(N 97)], trans = 4},
{fin = [(N 95)], trans = 5},
{fin = [(N 95)], trans = 6},
{fin = [(N 95)], trans = 7},
{fin = [(N 95)], trans = 8},
{fin = [(N 62),(N 95)], trans = 5},
{fin = [(N 95),(N 97)], trans = 5},
{fin = [(N 95),(N 97)], trans = 11},
{fin = [(N 95)], trans = 12},
{fin = [(N 95)], trans = 13},
{fin = [(N 95)], trans = 14},
{fin = [(N 39),(N 95)], trans = 5},
{fin = [(N 95),(N 97)], trans = 16},
{fin = [(N 95)], trans = 17},
{fin = [(N 57),(N 95)], trans = 5},
{fin = [(N 95),(N 97)], trans = 19},
{fin = [(N 95)], trans = 20},
{fin = [(N 95)], trans = 21},
{fin = [(N 95)], trans = 22},
{fin = [(N 95)], trans = 23},
{fin = [(N 69),(N 95)], trans = 5},
{fin = [(N 90),(N 97)], trans = 25},
{fin = [(N 90)], trans = 25},
{fin = [(N 90)], trans = 27},
{fin = [(N 90),(N 97)], trans = 28},
{fin = [(N 90)], trans = 29},
{fin = [(N 90)], trans = 30},
{fin = [(N 90)], trans = 31},
{fin = [(N 90)], trans = 32},
{fin = [(N 90)], trans = 33},
{fin = [(N 90)], trans = 34},
{fin = [(N 90)], trans = 35},
{fin = [(N 90)], trans = 36},
{fin = [(N 33),(N 90)], trans = 25},
{fin = [(N 97)], trans = 38},
{fin = [(N 53)], trans = 0},
{fin = [(N 41),(N 97)], trans = 0},
{fin = [(N 72),(N 97)], trans = 41},
{fin = [(N 72)], trans = 41},
{fin = [(N 97)], trans = 43},
{fin = [], trans = 44},
{fin = [], trans = 45},
{fin = [], trans = 46},
{fin = [], trans = 47},
{fin = [], trans = 48},
{fin = [(N 20)], trans = 0},
{fin = [], trans = 50},
{fin = [(N 20)], trans = 48},
{fin = [], trans = 44},
{fin = [(N 22),(N 97)], trans = 0},
{fin = [(N 50),(N 97)], trans = 54},
{fin = [(N 48)], trans = 0},
{fin = [(N 45),(N 97)], trans = 0},
{fin = [(N 43),(N 97)], trans = 0},
{fin = [(N 97)], trans = 58},
{fin = [], trans = 58},
{fin = [], trans = 60},
{fin = [(N 85)], trans = 0},
{fin = [(N 97)], trans = 62},
{fin = [], trans = 62},
{fin = [(N 8)], trans = 0},
{fin = [(N 4),(N 97)], trans = 65},
{fin = [(N 4)], trans = 65},
{fin = [(N 1)], trans = 0}])
end
structure StartStates =
struct
datatype yystartstate = STARTSTATE of int
(* start state definitions *)
val INITIAL = STARTSTATE 1;
end
type result = UserDeclarations.lexresult
exception LexerError (* raised if illegal leaf action tried *)
end
fun makeLexer yyinput =
let val yygone0=1
val yyb = Unsynchronized.ref "\n" (* buffer *)
val yybl = Unsynchronized.ref 1 (*buffer length *)
val yybufpos = Unsynchronized.ref 1 (* location of next character to use *)
val yygone = Unsynchronized.ref yygone0 (* position in file of beginning of buffer *)
val yydone = Unsynchronized.ref false (* eof found yet? *)
val yybegin = Unsynchronized.ref 1 (*Current 'start state' for lexer *)
val YYBEGIN = fn (Internal.StartStates.STARTSTATE x) =>
yybegin := x
fun lex () : Internal.result =
let fun continue() = lex() in
let fun scan (s,AcceptingLeaves : Internal.yyfinstate list list,l,i0) =
let fun action (i,nil) = raise LexError
| action (i,nil::l) = action (i-1,l)
| action (i,(node::acts)::l) =
case node of
Internal.N yyk =>
(let fun yymktext() = String.substring(!yyb,i0,i-i0)
val yypos = i0+ !yygone
open UserDeclarations Internal.StartStates
in (yybufpos := i; case yyk of
(* Application actions *)
1 => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex())
| 20 => (lex())
| 22 => let val yytext=yymktext() in Tokens.COMMA(yytext,inputPos_half yypos,inputPos_half yypos) end
| 33 => let val yytext=yymktext() in Tokens.FIXEDPOINT(yytext,inputPos_half yypos,inputPos_half yypos) end
| 39 => let val yytext=yymktext() in Tokens.WHERE(yytext,inputPos_half yypos,inputPos_half yypos) end
| 4 => (pos := (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))); lex())
| 41 => let val yytext=yymktext() in Tokens.COLON(yytext,inputPos_half yypos,inputPos_half yypos) end
| 43 => let val yytext=yymktext() in Tokens.PAREN_OPEN(yytext,inputPos_half yypos,inputPos_half yypos) end
| 45 => let val yytext=yymktext() in Tokens.PAREN_CLOSE(yytext,inputPos_half yypos,inputPos_half yypos) end
| 48 => let val yytext=yymktext() in Tokens.DOUBLE_ASTERISK(yytext,inputPos_half yypos,inputPos_half yypos) end
| 50 => let val yytext=yymktext() in Tokens.ASTERISK(yytext,inputPos_half yypos,inputPos_half yypos) end
| 53 => let val yytext=yymktext() in Tokens.DOUBLE_RARROW(yytext,inputPos_half yypos,inputPos_half yypos) end
| 57 => let val yytext=yymktext() in Tokens.ONE(yytext,inputPos_half yypos,inputPos_half yypos) end
| 62 => let val yytext=yymktext() in Tokens.ZERO(yytext,inputPos_half yypos,inputPos_half yypos) end
| 69 => let val yytext=yymktext() in Tokens.ATTACK(yytext,inputPos_half yypos,inputPos_half yypos) end
| 72 => let val yytext=yymktext() in Tokens.INTEGER_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos) end
| 8 => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex())
| 85 => let val yytext=yymktext() in Tokens.STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos) end
| 90 => let val yytext=yymktext() in Tokens.UPPER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos) end
| 95 => let val yytext=yymktext() in Tokens.LOWER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos) end
| 97 => let val yytext=yymktext() in error ("ignoring bad character "^yytext,
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))),
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))));
lex() end
| _ => raise Internal.LexerError
) end )
val {fin,trans} = Vector.sub(Internal.tab, s)
val NewAcceptingLeaves = fin::AcceptingLeaves
in if l = !yybl then
if trans = #trans(Vector.sub(Internal.tab,0))
then action(l,NewAcceptingLeaves
) else let val newchars= if !yydone then "" else yyinput 1024
in if (String.size newchars)=0
then (yydone := true;
if (l=i0) then UserDeclarations.eof ()
else action(l,NewAcceptingLeaves))
else (if i0=l then yyb := newchars
else yyb := String.substring(!yyb,i0,l-i0)^newchars;
yygone := !yygone+i0;
yybl := String.size (!yyb);
scan (s,AcceptingLeaves,l-i0,0))
end
else let val NewChar = Char.ord(CharVector.sub(!yyb,l))
val NewChar = if NewChar<128 then NewChar else 128
val NewState = Char.ord(CharVector.sub(trans,NewChar))
in if NewState=0 then action(l,NewAcceptingLeaves)
else scan(NewState,NewAcceptingLeaves,l+1,i0)
end
end
(*
val start= if String.substring(!yyb,!yybufpos-1,1)="\n"
then !yybegin+1 else !yybegin
*)
in scan(!yybegin (* start *),nil,!yybufpos,!yybufpos)
end
end
in lex
end
end

View File

@ -0,0 +1,287 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
open Trac_Term
exception NotYetSupported of string
%%
%verbose
%eop EOF
%left
%name TracTransaction
%term EOF
| OPENP of string
| CLOSEP of string
| OPENB of string
| CLOSEB of string
| OPENSCRYPT of string
| CLOSESCRYPT of string
| COLON of string
| SEMICOLON of string
| SECCH of string
| AUTHCH of string
| CONFCH of string
| INSECCH of string
| FAUTHCH of string
| FSECCH of string
| PERCENT of string
| UNEQUAL of string
| EXCLAM of string
| DOT of string
| COMMA of string
| OPENSQB of string
| CLOSESQB of string
| UNION of string
| PROTOCOL of string
| KNOWLEDGE of string
| WHERE of string
| ACTIONS of string
| ABSTRACTION of string
| GOALS of string
| AUTHENTICATES of string
| WEAKLY of string
| ON of string
| TSECRET of string
| TBETWEEN of string
| Sets of string
| FUNCTIONS of string
| PUBLIC of string
| PRIVATE of string
| RECEIVE of string
| SEND of string
| IN of string
| NOTIN of string
| INSERT of string
| DELETE of string
| NEW of string
| ATTACK of string
| slash of string
| QUESTION of string
| equal of string
| TYPES of string
| SETS of string
| ARROW of string
| ANALYSIS of string
| TRANSACTIONS of string
| STRING_LITERAL of string
| UPPER_STRING_LITERAL of string
| LOWER_STRING_LITERAL of string
| UNDERSCORE of string
| INTEGER_LITERAL of string
| STAR of string
| OF of string
%nonterm START of TracProtocol.protocol
| name of string
| arity of string
| uident of string
| lident of string
| ident of string
| trac_protocol of TracProtocol.protocol
| protocol_spec of TracProtocol.protocol
| type_union of (string list)
| type_spec of (string * TracProtocol.type_spec_elem)
| type_specs of (string * TracProtocol.type_spec_elem) list
| idents of string list
| uidents of string list
| lidents of string list
| set_specs of TracProtocol.set_spec list
| set_spec of TracProtocol.set_spec
| priv_or_pub_fun_spec of TracProtocol.fun_spec
| fun_specs of TracProtocol.funT list
| fun_spec of TracProtocol.funT
| priv_fun_spec of TracProtocol.funT list
| pub_fun_spec of TracProtocol.funT list
| analysis_spec of TracProtocol.anaT
| transaction_spec_head of string option
| transaction_spec of TracProtocol.transaction list
| rule of TracProtocol.ruleT
| head of string * string list
| head_params of string list
| keys of Trac_Term.Msg list
| result of string list
| msg of Trac_Term.Msg
| msgs of Trac_Term.Msg list
| setexp of string * Trac_Term.Msg list
| action of TracProtocol.prot_label * TracProtocol.action
| actions of (TracProtocol.prot_label * TracProtocol.action) list
| ineq_aux of string
| ineq of string * string
| ineqs of (string * string) list
| transaction of TracProtocol.transaction_name
| typ of string
| parameter of string * string
| parameters of (string * string) list
%pos (int * int * int)
%noshift EOF
%%
START: trac_protocol (trac_protocol)
trac_protocol: PROTOCOL COLON name protocol_spec (TracProtocol.update_name protocol_spec name)
protocol_spec: TYPES COLON type_specs protocol_spec (TracProtocol.update_type_spec protocol_spec type_specs)
| SETS COLON set_specs protocol_spec (TracProtocol.update_sets protocol_spec set_specs)
| FUNCTIONS COLON priv_or_pub_fun_spec protocol_spec (TracProtocol.update_functions protocol_spec (SOME priv_or_pub_fun_spec))
| ANALYSIS COLON analysis_spec protocol_spec (TracProtocol.update_analysis protocol_spec analysis_spec)
| transaction_spec_head COLON transaction_spec protocol_spec (TracProtocol.update_transactions transaction_spec_head protocol_spec transaction_spec)
| (TracProtocol.empty)
type_union: ident ([ident])
| ident UNION type_union (ident::type_union)
type_specs: type_spec ([type_spec])
| type_spec type_specs (type_spec::type_specs)
type_spec: ident equal OPENB lidents CLOSEB ((ident, TracProtocol.Consts lidents))
| ident equal type_union ((ident, TracProtocol.Union type_union))
set_specs: set_spec ([set_spec])
| set_spec set_specs (set_spec::set_specs)
set_spec: ident slash arity ((ident, arity))
priv_or_pub_fun_spec: pub_fun_spec priv_or_pub_fun_spec (TracProtocol.update_fun_public priv_or_pub_fun_spec pub_fun_spec)
| priv_fun_spec priv_or_pub_fun_spec (TracProtocol.update_fun_private priv_or_pub_fun_spec priv_fun_spec)
| (TracProtocol.fun_empty)
pub_fun_spec: PUBLIC fun_specs (fun_specs)
priv_fun_spec: PRIVATE fun_specs (fun_specs)
fun_specs: fun_spec ([fun_spec])
| fun_spec fun_specs (fun_spec::fun_specs)
fun_spec: ident slash arity ((ident, arity))
analysis_spec: rule ([rule])
| rule analysis_spec (rule::analysis_spec)
rule: head ARROW result ((head,[],result))
| head QUESTION keys ARROW result ((head,keys,result))
head: LOWER_STRING_LITERAL OPENP head_params CLOSEP ((LOWER_STRING_LITERAL,head_params))
head_params: UPPER_STRING_LITERAL ([UPPER_STRING_LITERAL])
| UPPER_STRING_LITERAL COMMA head_params ([UPPER_STRING_LITERAL]@head_params)
keys: msgs (msgs)
result: UPPER_STRING_LITERAL ([UPPER_STRING_LITERAL])
| UPPER_STRING_LITERAL COMMA result ([UPPER_STRING_LITERAL]@result)
transaction_spec_head: TRANSACTIONS (NONE)
| TRANSACTIONS OF LOWER_STRING_LITERAL (SOME LOWER_STRING_LITERAL)
transaction_spec: transaction actions DOT ([TracProtocol.mkTransaction transaction actions])
| transaction actions DOT transaction_spec ((TracProtocol.mkTransaction transaction actions)::transaction_spec)
ineq_aux: UNEQUAL UPPER_STRING_LITERAL (UPPER_STRING_LITERAL)
ineq: UPPER_STRING_LITERAL ineq_aux ((UPPER_STRING_LITERAL,ineq_aux))
ineqs: ineq ([ineq])
| ineq COMMA ineqs ([ineq]@ineqs)
transaction: ident OPENP parameters CLOSEP WHERE ineqs ((ident,parameters,ineqs))
| ident OPENP parameters CLOSEP ((ident,parameters,[]))
| ident OPENP CLOSEP ((ident,[],[]))
parameters: parameter ([parameter])
| parameter COMMA parameters (parameter::parameters)
parameter: ident COLON typ ((ident, typ))
typ: UPPER_STRING_LITERAL (UPPER_STRING_LITERAL)
| LOWER_STRING_LITERAL (LOWER_STRING_LITERAL)
actions: action ([action])
| action actions (action::actions)
action: RECEIVE msg ((TracProtocol.LabelN,TracProtocol.RECEIVE(msg)))
| SEND msg ((TracProtocol.LabelN,TracProtocol.SEND(msg)))
| msg IN setexp ((TracProtocol.LabelN,TracProtocol.IN(msg,setexp)))
| msg NOTIN setexp ((TracProtocol.LabelN,TracProtocol.NOTIN(msg,setexp)))
| msg NOTIN lident OPENP UNDERSCORE CLOSEP ((TracProtocol.LabelN,TracProtocol.NOTINANY(msg,lident)))
| INSERT msg setexp ((TracProtocol.LabelN,TracProtocol.INSERT(msg,setexp)))
| DELETE msg setexp ((TracProtocol.LabelN,TracProtocol.DELETE(msg,setexp)))
| NEW uident ((TracProtocol.LabelS,TracProtocol.NEW(uident)))
| ATTACK ((TracProtocol.LabelN,TracProtocol.ATTACK))
| STAR RECEIVE msg ((TracProtocol.LabelS,TracProtocol.RECEIVE(msg)))
| STAR SEND msg ((TracProtocol.LabelS,TracProtocol.SEND(msg)))
| STAR msg IN setexp ((TracProtocol.LabelS,TracProtocol.IN(msg,setexp)))
| STAR msg NOTIN setexp ((TracProtocol.LabelS,TracProtocol.NOTIN(msg,setexp)))
| STAR msg NOTIN lident OPENP UNDERSCORE CLOSEP ((TracProtocol.LabelS,TracProtocol.NOTINANY(msg,lident)))
| STAR INSERT msg setexp ((TracProtocol.LabelS,TracProtocol.INSERT(msg,setexp)))
| STAR DELETE msg setexp ((TracProtocol.LabelS,TracProtocol.DELETE(msg,setexp)))
setexp: lident ((lident,[]))
| lident OPENP msgs CLOSEP ((lident,msgs))
msg: uident (Var uident)
| lident (Const lident)
| lident OPENP msgs CLOSEP (Fun (lident,msgs))
msgs: msg ([msg])
| msg COMMA msgs (msg::msgs)
name: UPPER_STRING_LITERAL (UPPER_STRING_LITERAL)
| LOWER_STRING_LITERAL (LOWER_STRING_LITERAL)
uident: UPPER_STRING_LITERAL (UPPER_STRING_LITERAL)
uidents: uident ([uident])
| uident COMMA uidents (uident::uidents)
lident: LOWER_STRING_LITERAL (LOWER_STRING_LITERAL)
lidents: lident ([lident])
| lident COMMA lidents (lident::lidents)
ident: uident (uident)
| lident (lident)
idents: ident ([ident])
| ident COMMA idents (ident::idents)
arity: INTEGER_LITERAL (INTEGER_LITERAL)

View File

@ -0,0 +1,73 @@
signature TracTransaction_TOKENS =
sig
type ('a,'b) token
type svalue
val OF: (string) * 'a * 'a -> (svalue,'a) token
val STAR: (string) * 'a * 'a -> (svalue,'a) token
val INTEGER_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val UNDERSCORE: (string) * 'a * 'a -> (svalue,'a) token
val LOWER_STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val UPPER_STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val STRING_LITERAL: (string) * 'a * 'a -> (svalue,'a) token
val TRANSACTIONS: (string) * 'a * 'a -> (svalue,'a) token
val ANALYSIS: (string) * 'a * 'a -> (svalue,'a) token
val ARROW: (string) * 'a * 'a -> (svalue,'a) token
val SETS: (string) * 'a * 'a -> (svalue,'a) token
val TYPES: (string) * 'a * 'a -> (svalue,'a) token
val equal: (string) * 'a * 'a -> (svalue,'a) token
val QUESTION: (string) * 'a * 'a -> (svalue,'a) token
val slash: (string) * 'a * 'a -> (svalue,'a) token
val ATTACK: (string) * 'a * 'a -> (svalue,'a) token
val NEW: (string) * 'a * 'a -> (svalue,'a) token
val DELETE: (string) * 'a * 'a -> (svalue,'a) token
val INSERT: (string) * 'a * 'a -> (svalue,'a) token
val NOTIN: (string) * 'a * 'a -> (svalue,'a) token
val IN: (string) * 'a * 'a -> (svalue,'a) token
val SEND: (string) * 'a * 'a -> (svalue,'a) token
val RECEIVE: (string) * 'a * 'a -> (svalue,'a) token
val PRIVATE: (string) * 'a * 'a -> (svalue,'a) token
val PUBLIC: (string) * 'a * 'a -> (svalue,'a) token
val FUNCTIONS: (string) * 'a * 'a -> (svalue,'a) token
val Sets: (string) * 'a * 'a -> (svalue,'a) token
val TBETWEEN: (string) * 'a * 'a -> (svalue,'a) token
val TSECRET: (string) * 'a * 'a -> (svalue,'a) token
val ON: (string) * 'a * 'a -> (svalue,'a) token
val WEAKLY: (string) * 'a * 'a -> (svalue,'a) token
val AUTHENTICATES: (string) * 'a * 'a -> (svalue,'a) token
val GOALS: (string) * 'a * 'a -> (svalue,'a) token
val ABSTRACTION: (string) * 'a * 'a -> (svalue,'a) token
val ACTIONS: (string) * 'a * 'a -> (svalue,'a) token
val WHERE: (string) * 'a * 'a -> (svalue,'a) token
val KNOWLEDGE: (string) * 'a * 'a -> (svalue,'a) token
val PROTOCOL: (string) * 'a * 'a -> (svalue,'a) token
val UNION: (string) * 'a * 'a -> (svalue,'a) token
val CLOSESQB: (string) * 'a * 'a -> (svalue,'a) token
val OPENSQB: (string) * 'a * 'a -> (svalue,'a) token
val COMMA: (string) * 'a * 'a -> (svalue,'a) token
val DOT: (string) * 'a * 'a -> (svalue,'a) token
val EXCLAM: (string) * 'a * 'a -> (svalue,'a) token
val UNEQUAL: (string) * 'a * 'a -> (svalue,'a) token
val PERCENT: (string) * 'a * 'a -> (svalue,'a) token
val FSECCH: (string) * 'a * 'a -> (svalue,'a) token
val FAUTHCH: (string) * 'a * 'a -> (svalue,'a) token
val INSECCH: (string) * 'a * 'a -> (svalue,'a) token
val CONFCH: (string) * 'a * 'a -> (svalue,'a) token
val AUTHCH: (string) * 'a * 'a -> (svalue,'a) token
val SECCH: (string) * 'a * 'a -> (svalue,'a) token
val SEMICOLON: (string) * 'a * 'a -> (svalue,'a) token
val COLON: (string) * 'a * 'a -> (svalue,'a) token
val CLOSESCRYPT: (string) * 'a * 'a -> (svalue,'a) token
val OPENSCRYPT: (string) * 'a * 'a -> (svalue,'a) token
val CLOSEB: (string) * 'a * 'a -> (svalue,'a) token
val OPENB: (string) * 'a * 'a -> (svalue,'a) token
val CLOSEP: (string) * 'a * 'a -> (svalue,'a) token
val OPENP: (string) * 'a * 'a -> (svalue,'a) token
val EOF: 'a * 'a -> (svalue,'a) token
end
signature TracTransaction_LRVALS=
sig
structure Tokens : TracTransaction_TOKENS
structure ParserData:PARSER_DATA
sharing type ParserData.Token.token = Tokens.token
sharing type ParserData.svalue = Tokens.svalue
end

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,139 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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.
*)
structure Tokens = Tokens
open TracProtocol
type pos = int * int * int
type svalue = Tokens.svalue
type ('a,'b) token = ('a,'b) Tokens.token
type lexresult= (svalue,pos) token
val pos = ref (0,0,0)
fun eof () = Tokens.EOF((!pos,!pos))
fun error (e,p : (int * int * int),_) = TextIO.output (TextIO.stdOut,
String.concat[
"Line ", (Int.toString (#1 p)), "/",
(Int.toString (#2 p - #3 p)),": ", e, "\n"
])
fun inputPos yypos = ((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))),
(#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))))
fun inputPos_half yypos = (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))
%%
%header (functor TracTransactionLexFun(structure Tokens: TracTransaction_TOKENS));
alpha=[A-Za-z_];
upper=[A-Z];
lower=[a-z];
digit=[0-9];
ws = [\ \t];
%%
\n => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex());
{ws}+ => (pos := (#1 (!pos), yypos - (#3(!pos)), (#3 (!pos))); lex());
(#)[^\n]*\n => (pos := ((#1 (!pos)) + 1, yypos - (#3(!pos)),yypos ); lex());
"/*""/"*([^*/]|[^*]"/"|"*"[^/])*"*"*"*/" => (lex());
"(" => (Tokens.OPENP(yytext,inputPos_half yypos,inputPos_half yypos));
")" => (Tokens.CLOSEP(yytext,inputPos_half yypos,inputPos_half yypos));
"{" => (Tokens.OPENB(yytext,inputPos_half yypos,inputPos_half yypos));
"}" => (Tokens.CLOSEB(yytext,inputPos_half yypos,inputPos_half yypos));
"{|" => (Tokens.OPENSCRYPT(yytext,inputPos_half yypos,inputPos_half yypos));
"|}" => (Tokens.CLOSESCRYPT(yytext,inputPos_half yypos,inputPos_half yypos));
":" => (Tokens.COLON(yytext,inputPos_half yypos,inputPos_half yypos));
";" => (Tokens.SEMICOLON(yytext,inputPos_half yypos,inputPos_half yypos));
"->" => (Tokens.ARROW(yytext,inputPos_half yypos,inputPos_half yypos));
"%" => (Tokens.PERCENT(yytext,inputPos_half yypos,inputPos_half yypos));
"!=" => (Tokens.UNEQUAL(yytext,inputPos_half yypos,inputPos_half yypos));
"!" => (Tokens.EXCLAM (yytext,inputPos_half yypos,inputPos_half yypos));
"." => (Tokens.DOT(yytext,inputPos_half yypos,inputPos_half yypos));
"," => (Tokens.COMMA(yytext,inputPos_half yypos,inputPos_half yypos));
"[" => (Tokens.OPENSQB(yytext,inputPos_half yypos,inputPos_half yypos));
"]" => (Tokens.CLOSESQB(yytext,inputPos_half yypos,inputPos_half yypos));
"++" => (Tokens.UNION(yytext,inputPos_half yypos,inputPos_half yypos));
"Protocol" => (Tokens.PROTOCOL(yytext,inputPos_half yypos,inputPos_half yypos));
"Knowledge" => (Tokens.KNOWLEDGE(yytext,inputPos_half yypos,inputPos_half yypos));
"where" => (Tokens.WHERE(yytext,inputPos_half yypos,inputPos_half yypos));
"Types" => (Tokens.TYPES(yytext,inputPos_half yypos,inputPos_half yypos));
"Actions" => (Tokens.ACTIONS(yytext,inputPos_half yypos,inputPos_half yypos));
"Abstraction" => (Tokens.ABSTRACTION(yytext,inputPos_half yypos,inputPos_half yypos));
"Goals" => (Tokens.GOALS(yytext,inputPos_half yypos,inputPos_half yypos));
"authenticates" => (Tokens.AUTHENTICATES(yytext,inputPos_half yypos,inputPos_half yypos));
"weakly" => (Tokens.WEAKLY(yytext,inputPos_half yypos,inputPos_half yypos));
"on" => (Tokens.ON(yytext,inputPos_half yypos,inputPos_half yypos));
"secret" => (Tokens.TSECRET(yytext,inputPos_half yypos,inputPos_half yypos));
"between" => (Tokens.TBETWEEN(yytext,inputPos_half yypos,inputPos_half yypos));
"Sets" => (Tokens.SETS(yytext,inputPos_half yypos,inputPos_half yypos));
"Functions" => (Tokens.FUNCTIONS(yytext,inputPos_half yypos,inputPos_half yypos));
"Public" => (Tokens.PUBLIC(yytext,inputPos_half yypos,inputPos_half yypos));
"Private" => (Tokens.PRIVATE(yytext,inputPos_half yypos,inputPos_half yypos));
"Analysis" => (Tokens.ANALYSIS(yytext,inputPos_half yypos,inputPos_half yypos));
"Transactions" => (Tokens.TRANSACTIONS(yytext,inputPos_half yypos,inputPos_half yypos));
"receive" => (Tokens.RECEIVE(yytext,inputPos_half yypos,inputPos_half yypos));
"send" => (Tokens.SEND(yytext,inputPos_half yypos,inputPos_half yypos));
"in" => (Tokens.IN(yytext,inputPos_half yypos,inputPos_half yypos));
"notin" => (Tokens.NOTIN(yytext,inputPos_half yypos,inputPos_half yypos));
"insert" => (Tokens.INSERT(yytext,inputPos_half yypos,inputPos_half yypos));
"delete" => (Tokens.DELETE(yytext,inputPos_half yypos,inputPos_half yypos));
"new" => (Tokens.NEW(yytext,inputPos_half yypos,inputPos_half yypos));
"attack" => (Tokens.ATTACK(yytext,inputPos_half yypos,inputPos_half yypos));
"/" => (Tokens.slash(yytext,inputPos_half yypos,inputPos_half yypos));
"?" => (Tokens.QUESTION(yytext,inputPos_half yypos,inputPos_half yypos));
"=" => (Tokens.equal(yytext,inputPos_half yypos,inputPos_half yypos));
"_" => (Tokens.UNDERSCORE(yytext,inputPos_half yypos,inputPos_half yypos));
"*" => (Tokens.STAR(yytext,inputPos_half yypos,inputPos_half yypos));
"of" => (Tokens.OF(yytext,inputPos_half yypos,inputPos_half yypos));
{digit}+ => (Tokens.INTEGER_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
"'"({alpha}|{ws}|{digit})*(("."|"_"|"/"|"-")*({alpha}|{ws}|{digit})*)*"'" => (Tokens.STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
{lower}({alpha}|{digit})*("'")* => (Tokens.LOWER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
{upper}({alpha}|{digit})*("'")* => (Tokens.UPPER_STRING_LITERAL(yytext,inputPos_half yypos,inputPos_half yypos));
. => (error ("ignoring bad character "^yytext,
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))),
((#1 (!pos), yypos - (#3(!pos)), (#3 (!pos)))));
lex());

File diff suppressed because it is too large Load Diff

View File

@ -0,0 +1,118 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: trac_protocol_parser.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section \<open>Parser for the Trac Format\<close>
theory
trac_protocol_parser
imports
"trac_term"
begin
ML_file "trac_parser/trac_protocol.grm.sig"
ML_file "trac_parser/trac_protocol.lex.sml"
ML_file "trac_parser/trac_protocol.grm.sml"
ML\<open>
structure TracProtocolParser : sig
val parse_file: string -> TracProtocol.protocol
val parse_str: string -> TracProtocol.protocol
end =
struct
structure TracLrVals =
TracTransactionLrValsFun(structure Token = LrParser.Token)
structure TracLex =
TracTransactionLexFun(structure Tokens = TracLrVals.Tokens)
structure TracParser =
Join(structure LrParser = LrParser
structure ParserData = TracLrVals.ParserData
structure Lex = TracLex)
fun invoke lexstream =
let fun print_error (s,i:(int * int * int),_) =
error("Error, line .... " ^ (Int.toString (#1 i)) ^"."^(Int.toString (#2 i ))^ ", " ^ s ^ "\n")
in TracParser.parse(0,lexstream,print_error,())
end
fun parse_fp lexer = let
val dummyEOF = TracLrVals.Tokens.EOF((0,0,0),(0,0,0))
fun loop lexer =
let
val _ = (TracLex.UserDeclarations.pos := (0,0,0);())
val (res,lexer) = invoke lexer
val (nextToken,lexer) = TracParser.Stream.get lexer
in if TracParser.sameToken(nextToken,dummyEOF) then ((),res)
else loop lexer
end
in (#2(loop lexer))
end
fun parse_file tracFile =
let
val infile = TextIO.openIn tracFile
val lexer = TracParser.makeLexer (fn _ => case ((TextIO.inputLine) infile) of
SOME s => s
| NONE => "")
in
parse_fp lexer
handle LrParser.ParseError => TracProtocol.empty
end
fun parse_str str =
let
val parsed = Unsynchronized.ref false
fun input_string _ = if !parsed then "" else (parsed := true ;str)
val lexer = TracParser.makeLexer input_string
in
parse_fp lexer
handle LrParser.ParseError => TracProtocol.empty
end
end
\<close>
end

View File

@ -0,0 +1,565 @@
(*
(C) Copyright Andreas Viktor Hess, DTU, 2020
(C) Copyright Sebastian A. Mödersheim, DTU, 2020
(C) Copyright Achim D. Brucker, University of Exeter, 2020
(C) Copyright Anders Schlichtkrull, DTU, 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: trac_term.thy
Author: Andreas Viktor Hess, DTU
Author: Sebastian A. Mödersheim, DTU
Author: Achim D. Brucker, University of Exeter
Author: Anders Schlichtkrull, DTU
*)
section \<open>Abstract Syntax for Trac Terms\<close>
theory
trac_term
imports
"First_Order_Terms.Term"
"ml_yacc_lib"
(* Alternatively (provides, as a side-effect, ml-yacc-lib):
"HOL-TPTP.TPTP_Parser"
*)
begin
datatype cMsg = cVar "string * string"
| cConst string
| cFun "string * cMsg list"
ML\<open>
structure Trac_Utils =
struct
fun list_find p ts =
let
fun aux _ [] = NONE
| aux n (t::ts) =
if p t
then SOME (t,n)
else aux (n+1) ts
in
aux 0 ts
end
fun map_prod f (a,b) = (f a, f b)
fun list_product [] = [[]]
| list_product (xs::xss) =
List.concat (map (fn x => map (fn ys => x::ys) (list_product xss)) xs)
fun list_toString elem_toString xs =
let
fun aux [] = ""
| aux [x] = elem_toString x
| aux (x::y::xs) = elem_toString x ^ ", " ^ aux (y::xs)
in
"[" ^ aux xs ^ "]"
end
val list_to_str = list_toString (fn x => x)
fun list_triangle_product _ [] = []
| list_triangle_product f (x::xs) = map (f x) xs@list_triangle_product f xs
fun list_subseqs [] = [[]]
| list_subseqs (x::xs) = let val xss = list_subseqs xs in map (cons x) xss@xss end
fun list_intersect xs ys =
List.exists (fn x => member (op =) ys x) xs orelse
List.exists (fn y => member (op =) xs y) ys
fun list_partitions xs constrs =
let
val peq = eq_set (op =)
val pseq = eq_set peq
val psseq = eq_set pseq
fun illegal p q =
let
val pq = union (op =) p q
fun f (a,b) = member (op =) pq a andalso member (op =) pq b
in
List.exists f constrs
end
fun merges _ [] = []
| merges q (p::ps) =
if illegal p q then map (cons p) (merges q ps)
else (union (op =) p q::ps)::(map (cons p) (merges q ps))
fun merges_all [] = []
| merges_all (p::ps) = merges p ps@map (cons p) (merges_all ps)
fun step pss = fold (union pseq) (map merges_all pss) []
fun loop pss pssprev =
let val pss' = step pss
in if psseq (pss,pss') then pssprev else loop pss' (union pseq pss' pssprev)
end
val init = [map single xs]
in
loop init init
end
fun mk_unique [] = []
| mk_unique (x::xs) = x::mk_unique(List.filter (fn y => y <> x) xs)
fun list_rm_pair sel l x = filter (fn e => sel e <> x) l
fun list_minus list_rm l m = List.foldl (fn (a,b) => list_rm b a) l m
fun list_upto n =
let
fun aux m = if m >= n then [] else m::aux (m+1)
in
aux 0
end
end
\<close>
ML\<open>
structure Trac_Term (* : TRAC_TERM *) =
struct
open Trac_Utils
exception TypeError
type TypeDecl = string * string
datatype Msg = Var of string
| Const of string
| Fun of string * Msg list
| Attack
datatype VarType = EnumType of string
| ValueType
| Untyped
datatype cMsg = cVar of string * VarType
| cConst of string
| cFun of string * cMsg list
| cAttack
| cSet of string * cMsg list
| cAbs of (string * string list) list
| cOccursFact of cMsg
| cPrivFunSec
| cEnum of string
fun type_of et vt n =
case List.find (fn (v,_) => v = n) et of
SOME (_,t) => EnumType t
| NONE =>
if List.exists (fn v => v = n) vt
then ValueType
else Untyped
fun certifyMsg et vt (Var n) = cVar (n, type_of et vt n)
| certifyMsg _ _ (Const c) = cConst c
| certifyMsg et vt (Fun (f, ts)) = cFun (f, map (certifyMsg et vt) ts)
| certifyMsg _ _ Attack = cAttack
fun mk_Value_cVar x = cVar (x,ValueType)
val fv_Msg =
let
fun aux (Var x) = [x]
| aux (Fun (_,ts)) = List.concat (map aux ts)
| aux _ = []
in
mk_unique o aux
end
val fv_cMsg =
let
fun aux (cVar x) = [x]
| aux (cFun (_,ts)) = List.concat (map aux ts)
| aux (cSet (_,ts)) = List.concat (map aux ts)
| aux (cOccursFact bs) = aux bs
| aux _ = []
in
mk_unique o aux
end
fun subst_apply' (delta:(string * VarType) -> cMsg) (t:cMsg) =
case t of
cVar x => delta x
| cFun (f,ts) => cFun (f, map (subst_apply' delta) ts)
| cSet (s,ts) => cSet (s, map (subst_apply' delta) ts)
| cOccursFact bs => cOccursFact (subst_apply' delta bs)
| c => c
fun subst_apply (delta:(string * cMsg) list) =
subst_apply' (fn (n,tau) => (
case List.find (fn x => fst x = n) delta of
SOME x => snd x
| NONE => cVar (n,tau)))
end
\<close>
ML\<open>
structure TracProtocol (* : TRAC_TERM *) =
struct
open Trac_Utils
datatype type_spec_elem =
Consts of string list
| Union of string list
fun is_Consts t = case t of Consts _ => true | _ => false
fun the_Consts t = case t of Consts cs => cs | _ => error "Consts"
type type_spec = (string * type_spec_elem) list
type set_spec = (string * string)
fun extract_Consts (tspec:type_spec) =
(List.concat o map the_Consts o filter is_Consts o map snd) tspec
type funT = (string * string)
type fun_spec = {private: funT list, public: funT list}
type ruleT = (string * string list) * Trac_Term.Msg list * string list
type anaT = ruleT list
datatype prot_label = LabelN | LabelS
datatype action = RECEIVE of Trac_Term.Msg
| SEND of Trac_Term.Msg
| IN of Trac_Term.Msg * (string * Trac_Term.Msg list)
| NOTIN of Trac_Term.Msg * (string * Trac_Term.Msg list)
| NOTINANY of Trac_Term.Msg * string
| INSERT of Trac_Term.Msg * (string * Trac_Term.Msg list)
| DELETE of Trac_Term.Msg * (string * Trac_Term.Msg list)
| NEW of string
| ATTACK
datatype cAction = cReceive of Trac_Term.cMsg
| cSend of Trac_Term.cMsg
| cInequality of Trac_Term.cMsg * Trac_Term.cMsg
| cInSet of Trac_Term.cMsg * Trac_Term.cMsg
| cNotInSet of Trac_Term.cMsg * Trac_Term.cMsg
| cNotInAny of Trac_Term.cMsg * string
| cInsert of Trac_Term.cMsg * Trac_Term.cMsg
| cDelete of Trac_Term.cMsg * Trac_Term.cMsg
| cNew of string
| cAssertAttack
type transaction_name = string * (string * string) list * (string * string) list
type transaction={transaction:transaction_name,actions:(prot_label * action) list}
type cTransaction={
transaction:transaction_name,
receive_actions:(prot_label * cAction) list,
checksingle_actions:(prot_label * cAction) list,
checkall_actions:(prot_label * cAction) list,
fresh_actions:(prot_label * cAction) list,
update_actions:(prot_label * cAction) list,
send_actions:(prot_label * cAction) list,
attack_actions:(prot_label * cAction) list}
fun mkTransaction transaction actions = {transaction=transaction,
actions=actions}:transaction
fun is_RECEIVE a = case a of RECEIVE _ => true | _ => false
fun is_SEND a = case a of SEND _ => true | _ => false
fun is_IN a = case a of IN _ => true | _ => false
fun is_NOTIN a = case a of NOTIN _ => true | _ => false
fun is_NOTINANY a = case a of NOTINANY _ => true | _ => false
fun is_INSERT a = case a of INSERT _ => true | _ => false
fun is_DELETE a = case a of DELETE _ => true | _ => false
fun is_NEW a = case a of NEW _ => true | _ => false
fun is_ATTACK a = case a of ATTACK => true | _ => false
fun the_RECEIVE a = case a of RECEIVE t => t | _ => error "RECEIVE"
fun the_SEND a = case a of SEND t => t | _ => error "SEND"
fun the_IN a = case a of IN t => t | _ => error "IN"
fun the_NOTIN a = case a of NOTIN t => t | _ => error "NOTIN"
fun the_NOTINANY a = case a of NOTINANY t => t | _ => error "NOTINANY"
fun the_INSERT a = case a of INSERT t => t | _ => error "INSERT"
fun the_DELETE a = case a of DELETE t => t | _ => error "DELETE"
fun the_NEW a = case a of NEW t => t | _ => error "FRESH"
fun maybe_the_RECEIVE a = case a of RECEIVE t => SOME t | _ => NONE
fun maybe_the_SEND a = case a of SEND t => SOME t | _ => NONE
fun maybe_the_IN a = case a of IN t => SOME t | _ => NONE
fun maybe_the_NOTIN a = case a of NOTIN t => SOME t | _ => NONE
fun maybe_the_NOTINANY a = case a of NOTINANY t => SOME t | _ => NONE
fun maybe_the_INSERT a = case a of INSERT t => SOME t | _ => NONE
fun maybe_the_DELETE a = case a of DELETE t => SOME t | _ => NONE
fun maybe_the_NEW a = case a of NEW t => SOME t | _ => NONE
fun is_Receive a = case a of cReceive _ => true | _ => false
fun is_Send a = case a of cSend _ => true | _ => false
fun is_Inequality a = case a of cInequality _ => true | _ => false
fun is_InSet a = case a of cInSet _ => true | _ => false
fun is_NotInSet a = case a of cNotInSet _ => true | _ => false
fun is_NotInAny a = case a of cNotInAny _ => true | _ => false
fun is_Insert a = case a of cInsert _ => true | _ => false
fun is_Delete a = case a of cDelete _ => true | _ => false
fun is_Fresh a = case a of cNew _ => true | _ => false
fun is_Attack a = case a of cAssertAttack => true | _ => false
fun the_Receive a = case a of cReceive t => t | _ => error "Receive"
fun the_Send a = case a of cSend t => t | _ => error "Send"
fun the_Inequality a = case a of cInequality t => t | _ => error "Inequality"
fun the_InSet a = case a of cInSet t => t | _ => error "InSet"
fun the_NotInSet a = case a of cNotInSet t => t | _ => error "NotInSet"
fun the_NotInAny a = case a of cNotInAny t => t | _ => error "NotInAny"
fun the_Insert a = case a of cInsert t => t | _ => error "Insert"
fun the_Delete a = case a of cDelete t => t | _ => error "Delete"
fun the_Fresh a = case a of cNew t => t | _ => error "New"
fun maybe_the_Receive a = case a of cReceive t => SOME t | _ => NONE
fun maybe_the_Send a = case a of cSend t => SOME t | _ => NONE
fun maybe_the_Inequality a = case a of cInequality t => SOME t | _ => NONE
fun maybe_the_InSet a = case a of cInSet t => SOME t | _ => NONE
fun maybe_the_NotInSet a = case a of cNotInSet t => SOME t | _ => NONE
fun maybe_the_NotInAny a = case a of cNotInAny t => SOME t | _ => NONE
fun maybe_the_Insert a = case a of cInsert t => SOME t | _ => NONE
fun maybe_the_Delete a = case a of cDelete t => SOME t | _ => NONE
fun maybe_the_Fresh a = case a of cNew t => SOME t | _ => NONE
fun certifyAction et vt (lbl,SEND t) = (lbl,cSend (Trac_Term.certifyMsg et vt t))
| certifyAction et vt (lbl,RECEIVE t) = (lbl,cReceive (Trac_Term.certifyMsg et vt t))
| certifyAction et vt (lbl,IN (x,(s,ps))) = (lbl,cInSet
(Trac_Term.certifyMsg et vt x, Trac_Term.cSet (s, map (Trac_Term.certifyMsg et vt) ps)))
| certifyAction et vt (lbl,NOTIN (x,(s,ps))) = (lbl,cNotInSet
(Trac_Term.certifyMsg et vt x, Trac_Term.cSet (s, map (Trac_Term.certifyMsg et vt) ps)))
| certifyAction et vt (lbl,NOTINANY (x,s)) = (lbl,cNotInAny (Trac_Term.certifyMsg et vt x, s))
| certifyAction et vt (lbl,INSERT (x,(s,ps))) = (lbl,cInsert
(Trac_Term.certifyMsg et vt x, Trac_Term.cSet (s, map (Trac_Term.certifyMsg et vt) ps)))
| certifyAction et vt (lbl,DELETE (x,(s,ps))) = (lbl,cDelete
(Trac_Term.certifyMsg et vt x, Trac_Term.cSet (s, map (Trac_Term.certifyMsg et vt) ps)))
| certifyAction _ _ (lbl,NEW x) = (lbl,cNew x)
| certifyAction _ _ (lbl,ATTACK) = (lbl,cAssertAttack)
fun certifyTransaction (tr:transaction) =
let
val mk_cOccurs = Trac_Term.cOccursFact
fun mk_Value_cVar x = Trac_Term.cVar (x,Trac_Term.ValueType)
fun mk_cInequality x y = cInequality (mk_Value_cVar x, mk_Value_cVar y)
val mk_cInequalities = list_triangle_product mk_cInequality
val fresh_vals = map_filter (maybe_the_NEW o snd) (#actions tr)
val decl_vars = map fst (#2 (#transaction tr))
val neq_constrs = #3 (#transaction tr)
val _ = if List.exists (fn x => List.exists (fn y => x = y) fresh_vals) decl_vars
orelse List.exists (fn x => List.exists (fn y => x = y) decl_vars) fresh_vals
then error "the fresh and the declared variables must not overlap"
else ()
val _ = case List.find (fn (x,y) => x = y) neq_constrs of
SOME (x,y) => error ("illegal inequality constraint: " ^ x ^ " != " ^ y)
| NONE => ()
val nonfresh_vals = map fst (filter (fn x => snd x = "value") (#2 (#transaction tr)))
val enum_vars = filter (fn x => snd x <> "value") (#2 (#transaction tr))
fun lblS t = (LabelS,t)
val cactions = map (certifyAction enum_vars (nonfresh_vals@fresh_vals)) (#actions tr)
val nonfresh_occurs = map (lblS o cReceive o mk_cOccurs o mk_Value_cVar) nonfresh_vals
val receives = filter (is_Receive o snd) cactions
val value_inequalities = map lblS (mk_cInequalities nonfresh_vals)
val checksingles = filter (fn (_,a) => is_InSet a orelse is_NotInSet a) cactions
val checkalls = filter (is_NotInAny o snd) cactions
val updates = filter (fn (_,a) => is_Insert a orelse is_Delete a) cactions
val fresh = filter (is_Fresh o snd) cactions
val sends = filter (is_Send o snd) cactions
val fresh_occurs = map (lblS o cSend o mk_cOccurs o mk_Value_cVar) fresh_vals
val attack_signals = filter (is_Attack o snd) cactions
in
{transaction = #transaction tr,
receive_actions = nonfresh_occurs@receives,
checksingle_actions = value_inequalities@checksingles,
checkall_actions = checkalls,
fresh_actions = fresh,
update_actions = updates,
send_actions = sends@fresh_occurs,
attack_actions = attack_signals}:cTransaction
end
fun subst_apply_action (delta:(string * Trac_Term.cMsg) list) (lbl:prot_label,a:cAction) =
let
val apply = Trac_Term.subst_apply delta
in
case a of
cReceive t => (lbl,cReceive (apply t))
| cSend t => (lbl,cSend (apply t))
| cInequality (x,y) => (lbl,cInequality (apply x, apply y))
| cInSet (x,s) => (lbl,cInSet (apply x, apply s))
| cNotInSet (x,s) => (lbl,cNotInSet (apply x, apply s))
| cNotInAny (x,s) => (lbl,cNotInAny (apply x, s))
| cInsert (x,s) => (lbl,cInsert (apply x, apply s))
| cDelete (x,s) => (lbl,cDelete (apply x, apply s))
| cNew x => (lbl,cNew x)
| cAssertAttack => (lbl,cAssertAttack)
end
fun subst_apply_actions delta =
map (subst_apply_action delta)
type protocol = {
name:string
,type_spec:type_spec
,set_spec:set_spec list
,function_spec:fun_spec option
,analysis_spec:anaT
,transaction_spec:(string option * transaction list) list
,fixed_point: (Trac_Term.cMsg list * (string * string list) list list *
((string * string list) list * (string * string list) list) list) option
}
exception TypeError
val fun_empty = {
public=[]
,private=[]
}:fun_spec
fun update_fun_public (fun_spec:fun_spec) public =
({public = public
,private = #private fun_spec
}):fun_spec
fun update_fun_private (fun_spec:fun_spec) private =
({public = #public fun_spec
,private = private
}):fun_spec
val empty={
name=""
,type_spec=[]
,set_spec=[]
,function_spec=NONE
,analysis_spec=[]
,transaction_spec=[]
,fixed_point = NONE
}:protocol
fun update_name (protocol_spec:protocol) name =
({name = name
,type_spec = #type_spec protocol_spec
,set_spec = #set_spec protocol_spec
,function_spec = #function_spec protocol_spec
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_sets (protocol_spec:protocol) set_spec =
({name = #name protocol_spec
,type_spec = #type_spec protocol_spec
,set_spec =
if has_duplicates (op =) (map fst set_spec)
then error "Multiple declarations of the same set family"
else set_spec
,function_spec = #function_spec protocol_spec
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_type_spec (protocol_spec:protocol) type_spec =
({name = #name protocol_spec
,type_spec =
if has_duplicates (op =) (map fst type_spec)
then error "Multiple declarations of the same enumeration type"
else type_spec
,set_spec = #set_spec protocol_spec
,function_spec = #function_spec protocol_spec
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_functions (protocol_spec:protocol) function_spec =
({name = #name protocol_spec
,type_spec = #type_spec protocol_spec
,set_spec = #set_spec protocol_spec
,function_spec = case function_spec of
SOME fs =>
if has_duplicates (op =) (map fst ((#public fs)@(#private fs)))
then error "Multiple declarations of the same constant or function symbol"
else SOME fs
| NONE => NONE
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_analysis (protocol_spec:protocol) analysis_spec =
({name = #name protocol_spec
,type_spec = #type_spec protocol_spec
,set_spec = #set_spec protocol_spec
,function_spec = #function_spec protocol_spec
,analysis_spec =
if has_duplicates (op =) (map (#1 o #1) analysis_spec)
then error "Multiple analysis rules declared for the same function symbol"
else if List.exists (has_duplicates (op =)) (map (#2 o #1) analysis_spec)
then error "The heads of the analysis rules must be linear terms"
else if let fun f ((_,xs),ts,ys) =
subset (op =) (ys@List.concat (map Trac_Term.fv_Msg ts), xs)
in List.exists (not o f) analysis_spec end
then error "Variables occurring in the body of an analysis rule should also occur in its head"
else analysis_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_transactions (prot_name:string option) (protocol_spec:protocol) transaction_spec =
({name = #name protocol_spec
,type_spec = #type_spec protocol_spec
,set_spec = #set_spec protocol_spec
,function_spec = #function_spec protocol_spec
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = (prot_name,transaction_spec)::(#transaction_spec protocol_spec)
,fixed_point = #fixed_point protocol_spec
}):protocol
fun update_fixed_point (protocol_spec:protocol) fixed_point =
({name = #name protocol_spec
,type_spec = #type_spec protocol_spec
,set_spec = #set_spec protocol_spec
,function_spec = #function_spec protocol_spec
,analysis_spec = #analysis_spec protocol_spec
,transaction_spec = #transaction_spec protocol_spec
,fixed_point = fixed_point
}):protocol
end
\<close>
end

31
CITATION Normal file
View File

@ -0,0 +1,31 @@
To cite the use of this formal theory, please use
Andreas V. Hess, Sebastian Mödersheim, Achim D. Brucker, and Anders Schlichtkrull.
Automated Stateful Protocol Verification. In Archive of Formal Proofs, 2020.
http://www.isa-afp.org/entries/Automated_Stateful_Protocol_Verification.html,
Formal proof development
A BibTeX entry for LaTeX users is
Article{ hess.ea:automated:2020,
abstract= {In protocol verification we observe a wide spectrum from fully
automated methods to interactive theorem proving with proof
assistants like Isabelle/HOL. In this AFP entry, we present a
fully-automated approach for verifying stateful security protocols,
i.e., protocols with mutable state that may span several sessions.
The approach supports reachability goals like secrecy and
authentication. We also include a simple user-friendly
transaction-based protocol specification language that is embedded
into Isabelle.},
author = {Andreas V. Hess and Sebastian M{\"o}dersheim and Achim D. Brucker and Anders Schlichtkrull},
date = {2020-04-08},
file = {https://www.brucker.ch/bibliography/download/2020/hess.ea-automated-outline-2020.pdf},
filelabel= {Outline},
issn = {2150-914x},
journal = {Archive of Formal Proofs},
month = {apr},
note = {\url{http://www.isa-afp.org/entries/Automated_Stateful_Protocol_Verification.html}, Formal proof development},
pdf = {https://www.brucker.ch/bibliography/download/2020/hess.ea-automated-2020.pdf},
title = {Automated Stateful Protocol Verification},
url = {https://www.brucker.ch/bibliography/abstract/hess.ea-automated-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.

58
README.md Normal file
View File

@ -0,0 +1,58 @@
# Automated Stateful Protocol Verification
This git repository contains a local mirror of
[Stateful Protocol Composition and Typing](https://www.isa-afp.org/entries/Automated_Stateful_Protocol_Verification.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 Automated Stateful Protocol Verification
entry) at a later stage.
## Installation
This project depends on another [AFP](https://www.isa-afp.org) entry:
[Stateful Protocol Composition and Typing](https://www.isa-afp.org/entries/Stateful_Protocol_Composition_and_Typing.html).
Please follow the [official guidelines](https://www.isa-afp.org/using.shtml)
for installing the AFP locally. For short:
* [Download](https://www.isa-afp.org/release/afp-current.tar.gz) the complete AFP
* Extract the downloaded archive to an directory of your choice
* Let's assume the extracted archive lives in `/home/isabelle/afp`, now execute:
```console
achim@logicalhacking:~$ echo "/home/isabelle/afp/thys" >> ~/.isabelle/Isabelle2020/ROOTS
```
## How to build
```console
achim@logicalhacking:~$ isabelle build -D Automated_Stateful_Protocol_Verification
```
## Authors
* Andreas V. Hess
* [Sebastian Mödersheim](https://people.compute.dtu.dk/samo/)
* [Achim D. Brucker](http://www.brucker.ch/)
* [Anders Schlichtkrull](https://people.compute.dtu.dk/andschl/)
## 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/Automated_Stateful_Protocol_Verification>.
## Publications
* Andreas V. Hess, Sebastian Mödersheim, Achim D. Brucker, and Anders
Schlichtkrull. Automated Stateful Protocol Verification. In Archive
of Formal Proofs, 2020.
<http://www.isa-afp.org/entries/Automated_Stateful_Protocol_Verification.html>,
Formal proof development