diff -r 8fb53badad99 -r cdea44c775fa src/HOL/Metis_Examples/Message.thy --- a/src/HOL/Metis_Examples/Message.thy Wed Dec 30 18:07:10 2015 +0100 +++ b/src/HOL/Metis_Examples/Message.thy Wed Dec 30 18:17:28 2015 +0100 @@ -47,21 +47,17 @@ | Crypt key msg --{*Encryption, public- or shared-key*} -text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*} +text{*Concrete syntax: messages appear as \\A,B,NA\\, etc...*} syntax - "_MTuple" :: "['a, args] => 'a * 'b" ("(2{|_,/ _|})") - -syntax (xsymbols) "_MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") - translations - "{|x, y, z|}" == "{|x, {|y, z|}|}" - "{|x, y|}" == "CONST MPair x y" + "\x, y, z\" == "\x, \y, z\\" + "\x, y\" == "CONST MPair x y" definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where --{*Message Y paired with a MAC computed with the help of X*} - "Hash[X] Y == {| Hash{|X,Y|}, Y|}" + "Hash[X] Y == \ Hash\X,Y\, Y\" definition keysFor :: "msg set => key set" where --{*Keys useful to decrypt elements of a message set*} @@ -75,8 +71,8 @@ for H :: "msg set" where Inj [intro]: "X \ H ==> X \ parts H" - | Fst: "{|X,Y|} \ parts H ==> X \ parts H" - | Snd: "{|X,Y|} \ parts H ==> Y \ parts H" + | Fst: "\X,Y\ \ parts H ==> X \ parts H" + | Snd: "\X,Y\ \ parts H ==> Y \ parts H" | Body: "Crypt K X \ parts H ==> X \ parts H" lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" @@ -134,7 +130,7 @@ lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H" by (unfold keysFor_def, auto) -lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H" +lemma keysFor_insert_MPair [simp]: "keysFor (insert \X,Y\ H) = keysFor H" by (unfold keysFor_def, auto) lemma keysFor_insert_Crypt [simp]: @@ -151,7 +147,7 @@ subsection{*Inductive relation "parts"*} lemma MPair_parts: - "[| {|X,Y|} \ parts H; + "[| \X,Y\ \ parts H; [| X \ parts H; Y \ parts H |] ==> P |] ==> P" by (blast dest: parts.Fst parts.Snd) @@ -294,8 +290,8 @@ done lemma parts_insert_MPair [simp]: - "parts (insert {|X,Y|} H) = - insert {|X,Y|} (parts (insert X (insert Y H)))" + "parts (insert \X,Y\ H) = + insert \X,Y\ (parts (insert X (insert Y H)))" apply (rule equalityI) apply (rule subsetI) apply (erule parts.induct, auto) @@ -325,8 +321,8 @@ for H :: "msg set" where Inj [intro,simp] : "X \ H ==> X \ analz H" - | Fst: "{|X,Y|} \ analz H ==> X \ analz H" - | Snd: "{|X,Y|} \ analz H ==> Y \ analz H" + | Fst: "\X,Y\ \ analz H ==> X \ analz H" + | Snd: "\X,Y\ \ analz H ==> Y \ analz H" | Decrypt [dest]: "[|Crypt K X \ analz H; Key(invKey K): analz H|] ==> X \ analz H" @@ -340,7 +336,7 @@ text{*Making it safe speeds up proofs*} lemma MPair_analz [elim!]: - "[| {|X,Y|} \ analz H; + "[| \X,Y\ \ analz H; [| X \ analz H; Y \ analz H |] ==> P |] ==> P" by (blast dest: analz.Fst analz.Snd) @@ -424,8 +420,8 @@ done lemma analz_insert_MPair [simp]: - "analz (insert {|X,Y|} H) = - insert {|X,Y|} (analz (insert X (insert Y H)))" + "analz (insert \X,Y\ H) = + insert \X,Y\ (analz (insert X (insert Y H)))" apply (rule equalityI) apply (rule subsetI) apply (erule analz.induct, auto) @@ -541,7 +537,7 @@ text{*If there are no pairs or encryptions then analz does nothing*} lemma analz_trivial: - "[| \X Y. {|X,Y|} \ H; \X K. Crypt K X \ H |] ==> analz H = H" + "[| \X Y. \X,Y\ \ H; \X K. Crypt K X \ H |] ==> analz H = H" apply safe apply (erule analz.induct, blast+) done @@ -572,7 +568,7 @@ | Agent [intro]: "Agent agt \ synth H" | Number [intro]: "Number n \ synth H" | Hash [intro]: "X \ synth H ==> Hash X \ synth H" - | MPair [intro]: "[|X \ synth H; Y \ synth H|] ==> {|X,Y|} \ synth H" + | MPair [intro]: "[|X \ synth H; Y \ synth H|] ==> \X,Y\ \ synth H" | Crypt [intro]: "[|X \ synth H; Key(K) \ H|] ==> Crypt K X \ synth H" text{*Monotonicity*} @@ -584,7 +580,7 @@ inductive_cases Nonce_synth [elim!]: "Nonce n \ synth H" inductive_cases Key_synth [elim!]: "Key K \ synth H" inductive_cases Hash_synth [elim!]: "Hash X \ synth H" -inductive_cases MPair_synth [elim!]: "{|X,Y|} \ synth H" +inductive_cases MPair_synth [elim!]: "\X,Y\ \ synth H" inductive_cases Crypt_synth [elim!]: "Crypt K X \ synth H"