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