(***************************************************************************** * ESPL --- an embedded security protocol logic * http://people.inf.ethz.ch/meiersi/espl/ * * Copyright (c) 2009-2011, Simon Meier, ETH Zurich, Switzerland * * Extension to compromising adversaries: * * Copyright (c) 2010-2011, Martin Schaub, ETH Zurich, Switzerland * * All rights reserved. See file LICENCE for more information. ******************************************************************************) (* 6th Generation: Adapt Syntax to make it more readable, also in ASCII notation Adapted to paper. *) theory Syntax imports ExplicitModel begin section{* User Level Syntax *} text{* The syntax adaption is currently done in a separate file and AFTER all developments of the logic, because it is not fixed yet. *} subsubsection{* General Cryptographic Messages *} notation (xsymbols) ExecMessage.inv ("[_]\<^isup>-\<^isup>1" [60] 58) subsubsection{* Specification Messages *} abbreviation "sC" where "sC x \ PConst x" abbreviation "sN" where "sN x \ PFresh x" abbreviation "sMV" where "sMV x \ PVar (MVar x)" abbreviation "sAV" where "sAV x \ PVar (AVar x)" text{* Note the following abbreviations are kept for backwards compatibility. They are deprecated for new developments. *} abbreviation "sLC" where "sLC x \ PConst x" abbreviation "sLN" where "sLN x \ PFresh x" abbreviation "sLMV" where "sLMV x \ PVar (MVar x)" abbreviation "sLAV" where "sLAV x \ PVar (AVar x)" abbreviation "sPK" where "sPK x \ PAsymPK (sAV x)" abbreviation "sSK" where "sSK x \ PAsymSK (sAV x)" abbreviation "sK" where "sK x y \ PSymK (sAV x) (sAV y)" abbreviation "sShrK" where "sShrK A \ PShrK A" syntax "@PTuple" :: "['a, args] => 'a * 'b" ("(2<|_,/ _|>)") syntax (xsymbols) "@PTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") translations "<|x, y, z|>" == "<|x, <|y, z|> |>" "<|x, y|>" == "CONST PTup x y" subsubsection{* Execution Messages *} abbreviation "C" where "C \ EConst" abbreviation "N" where "N \ ENonce" abbreviation "MV" where "MV x i \ (MVar x, i)" abbreviation "AV" where "AV x i \ (AVar x, i)" abbreviation "LC" where "LC x \ Lit (C x)" abbreviation "LN" where "LN x i \ Lit (N x i)" abbreviation "LAg" where "LAg x \ Lit (EAgent x)" end