# HG changeset patch # User wenzelm # Date 1518696276 -3600 # Node ID 560fbd6bc0474b6ed75b685ecd11b0ff556d371b # Parent e4e57da0583a28808e0d84fc6c6ae331d1a91128# Parent ce654b0e6d69c3afb4b6c20470807eb5e9647b3e merged diff -r e4e57da0583a -r 560fbd6bc047 src/Benchmarks/Quickcheck_Benchmark/Needham_Schroeder_Base.thy --- a/src/Benchmarks/Quickcheck_Benchmark/Needham_Schroeder_Base.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Benchmarks/Quickcheck_Benchmark/Needham_Schroeder_Base.thy Thu Feb 15 13:04:36 2018 +0100 @@ -49,7 +49,7 @@ | 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" + "[|Crypt K X \ analz H; Key(invKey K) \ analz H|] ==> X \ analz H" inductive_set synth :: "msg set => msg set" @@ -168,7 +168,7 @@ lemma [code]: "analz H = (let - H' = H \ (\((%m. case m of \X, Y\ => {X, Y} | Crypt K X => if Key (invKey K) : H then {X} else {} | _ => {}) ` H)) + H' = H \ (\((%m. case m of \X, Y\ => {X, Y} | Crypt K X => if Key (invKey K) \ H then {X} else {} | _ => {}) ` H)) in if H' = H then H else analz H')" sorry @@ -180,7 +180,7 @@ definition synth' :: "msg set => msg => bool" where - "synth' H m = (m : synth H)" + "synth' H m = (m \ synth H)" lemmas [code_pred_intro] = synth.intros[folded synth'_def] diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Prog_Prove/Isar.thy --- a/src/Doc/Prog_Prove/Isar.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Prog_Prove/Isar.thy Thu Feb 15 13:04:36 2018 +0100 @@ -332,7 +332,7 @@ \begin{minipage}[t]{.4\textwidth} \isa{% \ -(*<*)lemma "ALL x. P x" proof-(*>*) +(*<*)lemma "\x. P x" proof-(*>*) show "\x. P(x)" proof fix x @@ -346,7 +346,7 @@ \begin{minipage}[t]{.4\textwidth} \isa{% \ -(*<*)lemma "EX x. P(x)" proof-(*>*) +(*<*)lemma "\x. P(x)" proof-(*>*) show "\x. P(x)" proof text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ @@ -370,7 +370,7 @@ How to reason forward from \noquotes{@{prop[source] "\x. P(x)"}}: \end{isamarkuptext}% \ -(*<*)lemma True proof- assume 1: "EX x. P x"(*>*) +(*<*)lemma True proof- assume 1: "\x. P x"(*>*) have "\x. P(x)" (*<*)by(rule 1)(*>*)text_raw\\ \isasymproof\\\ then obtain x where p: "P(x)" by blast (*<*)oops(*>*) @@ -1066,7 +1066,7 @@ reasoning backwards: by which rules could some given fact have been proved? For the inductive definition of @{const ev}, rule inversion can be summarized like this: -@{prop[display]"ev n \ n = 0 \ (EX k. n = Suc(Suc k) \ ev k)"} +@{prop[display]"ev n \ n = 0 \ (\k. n = Suc(Suc k) \ ev k)"} The realisation in Isabelle is a case analysis. A simple example is the proof that @{prop"ev n \ ev (n - 2)"}. We already went through the details informally in \autoref{sec:Logic:even}. This @@ -1223,7 +1223,7 @@ \begin{exercise} Define a recursive function @{text "elems ::"} @{typ"'a list \ 'a set"} -and prove @{prop "x : elems xs \ \ys zs. xs = ys @ x # zs \ x \ elems ys"}. +and prove @{prop "x \ elems xs \ \ys zs. xs = ys @ x # zs \ x \ elems ys"}. \end{exercise} \begin{exercise} diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Prog_Prove/Logic.thy --- a/src/Doc/Prog_Prove/Logic.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Prog_Prove/Logic.thy Thu Feb 15 13:04:36 2018 +0100 @@ -97,7 +97,7 @@ \noquotes{@{term[source] "{t | x y. P}"}}\index{$IMP042@@{text"{t |x. P}"}}, where @{text "x y"} are those free variables in @{text t} that occur in @{text P}. -This is just a shorthand for @{term"{v. EX x y. v = t \ P}"}, where +This is just a shorthand for @{term"{v. \x y. v = t \ P}"}, where @{text v} is a new variable. For example, @{term"{x+y|x. x \ A}"} is short for \noquotes{@{term[source]"{v. \x. v = x+y \ x \ A}"}}. \end{warn} @@ -111,8 +111,8 @@ @{text "\"} & \texttt{\char`\\\char`\} & \texttt{Int} \end{tabular} \end{center} -Sets also allow bounded quantifications @{prop"ALL x : A. P"} and -@{prop"EX x : A. P"}. +Sets also allow bounded quantifications @{prop"\x \ A. P"} and +@{prop"\x \ A. P"}. For the more ambitious, there are also @{text"\"}\index{$HOLSet6@\isasymUnion} and @{text"\"}\index{$HOLSet7@\isasymInter}: @@ -703,7 +703,7 @@ that maps a binary predicate to another binary predicate: if @{text r} is of type @{text"\ \ \ \ bool"} then @{term "star r"} is again of type @{text"\ \ \ \ bool"}, and @{prop"star r x y"} means that @{text x} and @{text y} are in -the relation @{term"star r"}. Think @{term"r^*"} when you see @{term"star +the relation @{term"star r"}. Think @{term"r\<^sup>*"} when you see @{term"star r"}, because @{text"star r"} is meant to be the reflexive transitive closure. That is, @{prop"star r x y"} is meant to be true if from @{text x} we can reach @{text y} in finitely many @{text r} steps. This concept is naturally diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Prog_Prove/Types_and_funs.thy --- a/src/Doc/Prog_Prove/Types_and_funs.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Prog_Prove/Types_and_funs.thy Thu Feb 15 13:04:36 2018 +0100 @@ -525,7 +525,7 @@ simplify to @{const True}. We can split case-expressions similarly. For @{text nat} the rule looks like this: -@{prop[display,margin=65,indent=4]"P(case e of 0 \ a | Suc n \ b n) = ((e = 0 \ P a) & (ALL n. e = Suc n \ P(b n)))"} +@{prop[display,margin=65,indent=4]"P(case e of 0 \ a | Suc n \ b n) = ((e = 0 \ P a) \ (\n. e = Suc n \ P(b n)))"} Case expressions are not split automatically by @{text simp}, but @{text simp} can be instructed to do so: \begin{quote} diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Sugar/Sugar.thy --- a/src/Doc/Sugar/Sugar.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Sugar/Sugar.thy Thu Feb 15 13:04:36 2018 +0100 @@ -55,7 +55,7 @@ \subsection{Logic} -The formula @{prop[source]"\(\x. P x)"} is typeset as @{prop"~(EX x. P x)"}. +The formula @{prop[source]"\(\x. P x)"} is typeset as @{prop"\(\x. P x)"}. The predefined constructs @{text"if"}, @{text"let"} and @{text"case"} are set in sans serif font to distinguish them from diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/CTL/CTL.thy --- a/src/Doc/Tutorial/CTL/CTL.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/CTL/CTL.thy Thu Feb 15 13:04:36 2018 +0100 @@ -153,7 +153,7 @@ done text\\noindent -We assume the negation of the conclusion and prove @{term"s : lfp(af A)"}. +We assume the negation of the conclusion and prove @{term"s \ lfp(af A)"}. Unfolding @{const lfp} once and simplifying with the definition of @{const af} finishes the proof. @@ -214,14 +214,14 @@ txt\\noindent After simplification, the base case boils down to @{subgoals[display,indent=0,margin=70,goals_limit=1]} -The conclusion looks exceedingly trivial: after all, @{term t} is chosen such that @{prop"(s,t):M"} +The conclusion looks exceedingly trivial: after all, @{term t} is chosen such that @{prop"(s,t)\M"} holds. However, we first have to show that such a @{term t} actually exists! This reasoning is embodied in the theorem @{thm[source]someI2_ex}: @{thm[display,eta_contract=false]someI2_ex} When we apply this theorem as an introduction rule, @{text"?P x"} becomes -@{prop"(s, x) : M & Q x"} and @{text"?Q x"} becomes @{prop"(s,x) : M"} and we have to prove -two subgoals: @{prop"EX a. (s, a) : M & Q a"}, which follows from the assumptions, and -@{prop"(s, x) : M & Q x ==> (s,x) : M"}, which is trivial. Thus it is not surprising that +@{prop"(s, x) \ M \ Q x"} and @{text"?Q x"} becomes @{prop"(s,x) \ M"} and we have to prove +two subgoals: @{prop"\a. (s, a) \ M \ Q a"}, which follows from the assumptions, and +@{prop"(s, x) \ M \ Q x \ (s,x) \ M"}, which is trivial. Thus it is not surprising that @{text fast} can prove the base case quickly: \ diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/CTL/PDL.thy --- a/src/Doc/Tutorial/CTL/PDL.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/CTL/PDL.thy Thu Feb 15 13:04:36 2018 +0100 @@ -129,7 +129,7 @@ forward direction. Fortunately the converse induction theorem @{thm[source]converse_rtrancl_induct} already exists: @{thm[display,margin=60]converse_rtrancl_induct[no_vars]} -It says that if @{prop"(a,b):r\<^sup>*"} and we know @{prop"P b"} then we can infer +It says that if @{prop"(a,b)\r\<^sup>*"} and we know @{prop"P b"} then we can infer @{prop"P a"} provided each step backwards from a predecessor @{term z} of @{term b} preserves @{term P}. \ @@ -176,7 +176,7 @@ \footnote{We cannot use the customary @{text EX}: it is reserved as the \textsc{ascii}-equivalent of @{text"\"}.} with the intended semantics -@{prop[display]"(s \ EN f) = (EX t. (s,t) : M & t \ f)"} +@{prop[display]"(s \ EN f) = (\t. (s,t) \ M \ t \ f)"} Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How? Show that the semantics for @{term EF} satisfies the following recursion equation: @@ -190,7 +190,7 @@ apply(auto simp add: EF_lemma) done -lemma aux: "s \ f = (s : mc f)" +lemma aux: "s \ f = (s \ mc f)" apply(simp add: main) done diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Ifexpr/Ifexpr.thy --- a/src/Doc/Tutorial/Ifexpr/Ifexpr.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Ifexpr/Ifexpr.thy Thu Feb 15 13:04:36 2018 +0100 @@ -201,7 +201,7 @@ (case b of CIF b => False | VIF x => True | IF x y z => False))" lemma [simp]: - "ALL t e. valif (normif2 b t e) env = valif (IF b t e) env" + "\t e. valif (normif2 b t e) env = valif (IF b t e) env" apply(induct b) by(auto) @@ -209,7 +209,7 @@ apply(induct b) by(auto) -lemma [simp]: "ALL t e. normal2 t & normal2 e --> normal2(normif2 b t e)" +lemma [simp]: "\t e. normal2 t & normal2 e --> normal2(normif2 b t e)" apply(induct b) by(auto) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Inductive/Mutual.thy --- a/src/Doc/Tutorial/Inductive/Mutual.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Inductive/Mutual.thy Thu Feb 15 13:04:36 2018 +0100 @@ -67,7 +67,7 @@ text\\noindent Everything works as before, except that you write \commdx{inductive} instead of \isacommand{inductive\_set} and -@{prop"evn n"} instead of @{prop"n : Even"}. +@{prop"evn n"} instead of @{prop"n \ Even"}. When defining an n-ary relation as a predicate, it is recommended to curry the predicate: its type should be \mbox{@{text"\\<^sub>1 \ \ \ \\<^sub>n \ bool"}} rather than diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Inductive/Star.thy --- a/src/Doc/Tutorial/Inductive/Star.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Inductive/Star.thy Thu Feb 15 13:04:36 2018 +0100 @@ -77,7 +77,7 @@ is what we want, it is merely due to the order in which the assumptions occur in the subgoal, which it is not good practice to rely on. As a result, @{text"?xb"} becomes @{term x}, @{text"?xa"} becomes -@{term y} and @{text"?P"} becomes @{term"%u v. (u,z) : r*"}, thus +@{term y} and @{text"?P"} becomes @{term"\u v. (u,z) \ r*"}, thus yielding the above subgoal. So what went wrong? When looking at the instantiation of @{text"?P"} we see that it does not @@ -85,7 +85,7 @@ goal, of the pair @{term"(x,y)"} only @{term x} appears also in the conclusion, but not @{term y}. Thus our induction statement is too general. Fortunately, it can easily be specialized: -transfer the additional premise @{prop"(y,z):r*"} into the conclusion:\ +transfer the additional premise @{prop"(y,z)\r*"} into the conclusion:\ (*<*)oops(*>*) lemma rtc_trans[rule_format]: "(x,y) \ r* \ (y,z) \ r* \ (x,z) \ r*" @@ -157,7 +157,7 @@ \begin{exercise}\label{ex:converse-rtc-step} Show that the converse of @{thm[source]rtc_step} also holds: -@{prop[display]"[| (x,y) : r*; (y,z) : r |] ==> (x,z) : r*"} +@{prop[display]"[| (x,y) \ r*; (y,z) \ r |] ==> (x,z) \ r*"} \end{exercise} \begin{exercise} Repeat the development of this section, but starting with a definition of @@ -166,7 +166,7 @@ \end{exercise} \ (*<*) -lemma rtc_step2[rule_format]: "(x,y) : r* \ (y,z) : r --> (x,z) : r*" +lemma rtc_step2[rule_format]: "(x,y) \ r* \ (y,z) \ r \ (x,z) \ r*" apply(erule rtc.induct) apply blast apply(blast intro: rtc_step) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Misc/Plus.thy --- a/src/Doc/Tutorial/Misc/Plus.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Misc/Plus.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,7 +10,7 @@ text\\noindent and prove\ (*<*) -lemma [simp]: "!m. add m n = m+n" +lemma [simp]: "\m. add m n = m+n" apply(induct_tac n) by(auto) (*>*) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Misc/Tree2.thy --- a/src/Doc/Tutorial/Misc/Tree2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Misc/Tree2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -14,7 +14,7 @@ text\\noindent and prove\ (*<*) -lemma [simp]: "!xs. flatten2 t xs = flatten t @ xs" +lemma [simp]: "\xs. flatten2 t xs = flatten t @ xs" apply(induct_tac t) by(auto) (*>*) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Misc/prime_def.thy --- a/src/Doc/Tutorial/Misc/prime_def.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Misc/prime_def.thy Thu Feb 15 13:04:36 2018 +0100 @@ -7,12 +7,12 @@ A common mistake when writing definitions is to introduce extra free variables on the right-hand side. Consider the following, flawed definition (where @{text"dvd"} means ``divides''): -@{term[display,quotes]"prime(p) == 1 < p & (m dvd p --> (m=1 | m=p))"} +@{term[display,quotes]"prime(p) \ 1 < p \ (m dvd p \ (m=1 \ m=p))"} \par\noindent\hangindent=0pt Isabelle rejects this ``definition'' because of the extra @{term"m"} on the right-hand side, which would introduce an inconsistency (why?). The correct version is -@{term[display,quotes]"prime(p) == 1 < p & (!m. m dvd p --> (m=1 | m=p))"} +@{term[display,quotes]"prime(p) \ 1 < p \ (\m. m dvd p \ (m=1 \ m=p))"} \end{warn} \ (*<*) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Protocol/Event.thy --- a/src/Doc/Tutorial/Protocol/Event.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Protocol/Event.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,7 +12,7 @@ theory Event imports Message begin consts (*Initial states of agents -- parameter of the construction*) - initState :: "agent => msg set" + initState :: "agent \ msg set" datatype event = Says agent agent msg @@ -26,28 +26,28 @@ text\The constant "spies" is retained for compatibility's sake\ primrec - knows :: "agent => event list => msg set" + knows :: "agent \ event list \ msg set" where knows_Nil: "knows A [] = initState A" | knows_Cons: "knows A (ev # evs) = (if A = Spy then (case ev of - Says A' B X => insert X (knows Spy evs) - | Gets A' X => knows Spy evs - | Notes A' X => + Says A' B X \ insert X (knows Spy evs) + | Gets A' X \ knows Spy evs + | Notes A' X \ if A' \ bad then insert X (knows Spy evs) else knows Spy evs) else (case ev of - Says A' B X => + Says A' B X \ if A'=A then insert X (knows A evs) else knows A evs - | Gets A' X => + | Gets A' X \ if A'=A then insert X (knows A evs) else knows A evs - | Notes A' X => + | Notes A' X \ if A'=A then insert X (knows A evs) else knows A evs))" abbreviation (input) - spies :: "event list => msg set" where + spies :: "event list \ msg set" where "spies == knows Spy" text\Spy has access to his own key for spoof messages, but Server is secure\ @@ -65,24 +65,24 @@ primrec (*Set of items that might be visible to somebody: complement of the set of fresh items*) - used :: "event list => msg set" + used :: "event list \ msg set" where used_Nil: "used [] = (UN B. parts (initState B))" | used_Cons: "used (ev # evs) = (case ev of - Says A B X => parts {X} \ used evs - | Gets A X => used evs - | Notes A X => parts {X} \ used evs)" + Says A B X \ parts {X} \ used evs + | Gets A X \ used evs + | Notes A X \ parts {X} \ used evs)" \ \The case for @{term Gets} seems anomalous, but @{term Gets} always follows @{term Says} in real protocols. Seems difficult to change. See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}.\ -lemma Notes_imp_used [rule_format]: "Notes A X \ set evs --> X \ used evs" +lemma Notes_imp_used [rule_format]: "Notes A X \ set evs \ X \ used evs" apply (induct_tac evs) apply (auto split: event.split) done -lemma Says_imp_used [rule_format]: "Says A B X \ set evs --> X \ used evs" +lemma Says_imp_used [rule_format]: "Says A B X \ set evs \ X \ used evs" apply (induct_tac evs) apply (auto split: event.split) done @@ -103,7 +103,7 @@ on whether @{term "A=Spy"} and whether @{term "A\bad"}\ lemma knows_Spy_Notes [simp]: "knows Spy (Notes A X # evs) = - (if A:bad then insert X (knows Spy evs) else knows Spy evs)" + (if A\bad then insert X (knows Spy evs) else knows Spy evs)" by simp lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" @@ -123,13 +123,13 @@ text\Spy sees what is sent on the traffic\ lemma Says_imp_knows_Spy [rule_format]: - "Says A B X \ set evs --> X \ knows Spy evs" + "Says A B X \ set evs \ X \ knows Spy evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done lemma Notes_imp_knows_Spy [rule_format]: - "Notes A X \ set evs --> A: bad --> X \ knows Spy evs" + "Notes A X \ set evs \ A \ bad \ X \ knows Spy evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done @@ -158,7 +158,7 @@ by simp lemma knows_Gets: - "A \ Spy --> knows A (Gets A X # evs) = insert X (knows A evs)" + "A \ Spy \ knows A (Gets A X # evs) = insert X (knows A evs)" by simp @@ -172,14 +172,14 @@ by (simp add: subset_insertI) text\Agents know what they say\ -lemma Says_imp_knows [rule_format]: "Says A B X \ set evs --> X \ knows A evs" +lemma Says_imp_knows [rule_format]: "Says A B X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) apply blast done text\Agents know what they note\ -lemma Notes_imp_knows [rule_format]: "Notes A X \ set evs --> X \ knows A evs" +lemma Notes_imp_knows [rule_format]: "Notes A X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) apply blast @@ -187,7 +187,7 @@ text\Agents know what they receive\ lemma Gets_imp_knows_agents [rule_format]: - "A \ Spy --> Gets A X \ set evs --> X \ knows A evs" + "A \ Spy \ Gets A X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done @@ -196,8 +196,8 @@ text\What agents DIFFERENT FROM Spy know was either said, or noted, or got, or known initially\ lemma knows_imp_Says_Gets_Notes_initState [rule_format]: - "[| X \ knows A evs; A \ Spy |] ==> EX B. - Says A B X \ set evs | Gets A X \ set evs | Notes A X \ set evs | X \ initState A" + "[| X \ knows A evs; A \ Spy |] ==> \B. + Says A B X \ set evs \ Gets A X \ set evs \ Notes A X \ set evs \ X \ initState A" apply (erule rev_mp) apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) @@ -207,8 +207,8 @@ text\What the Spy knows -- for the time being -- was either said or noted, or known initially\ lemma knows_Spy_imp_Says_Notes_initState [rule_format]: - "[| X \ knows Spy evs |] ==> EX A B. - Says A B X \ set evs | Notes A X \ set evs | X \ initState Spy" + "[| X \ knows Spy evs |] ==> \A B. + Says A B X \ set evs \ Notes A X \ set evs \ X \ initState Spy" apply (erule rev_mp) apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) @@ -222,7 +222,7 @@ lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] -lemma initState_into_used: "X \ parts (initState B) ==> X \ used evs" +lemma initState_into_used: "X \ parts (initState B) \ X \ used evs" apply (induct_tac "evs") apply (simp_all add: parts_insert_knows_A split: event.split, blast) done @@ -246,7 +246,7 @@ used_Nil [simp del] used_Cons [simp del] -text\For proving theorems of the form @{term "X \ analz (knows Spy evs) --> P"} +text\For proving theorems of the form @{term "X \ analz (knows Spy evs) \ P"} New events added by induction to "evs" are discarded. Provided this information isn't needed, the proof will be much shorter, since it will omit complicated reasoning about @{term analz}.\ @@ -286,7 +286,7 @@ method_setup analz_mono_contra = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\ - "for proving theorems of the form X \ analz (knows Spy evs) --> P" + "for proving theorems of the form X \ analz (knows Spy evs) \ P" subsubsection\Useful for case analysis on whether a hash is a spoof or not\ @@ -343,7 +343,7 @@ method_setup synth_analz_mono_contra = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\ - "for proving theorems of the form X \ synth (analz (knows Spy evs)) --> P" + "for proving theorems of the form X \ synth (analz (knows Spy evs)) \ P" (*>*) section\Event Traces \label{sec:events}\ diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Protocol/Message.thy Thu Feb 15 13:04:36 2018 +0100 @@ -39,7 +39,7 @@ specification (invKey) invKey [simp]: "invKey (invKey K) = K" - invKey_symmetric: "all_symmetric --> invKey = id" + invKey_symmetric: "all_symmetric \ invKey = id" by (rule exI [of _ id], auto) @@ -81,13 +81,13 @@ (*<*) text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ syntax - "_MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") + "_MTuple" :: "['a, args] \ 'a * 'b" ("(2\_,/ _\)") translations "\x, y, z\" == "\x, \y, z\\" "\x, y\" == "CONST MPair x y" -definition keysFor :: "msg set => key set" where +definition keysFor :: "msg set \ key set" where \ \Keys useful to decrypt elements of a message set\ "keysFor H == invKey ` {K. \X. Crypt K X \ H}" @@ -95,17 +95,17 @@ subsubsection\Inductive Definition of All Parts" of a Message\ inductive_set - parts :: "msg set => msg set" + parts :: "msg set \ msg set" 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" - | Body: "Crypt K X \ parts H ==> X \ parts H" + Inj [intro]: "X \ H \ X \ 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" text\Monotonicity\ -lemma parts_mono: "G \ H ==> parts(G) \ parts(H)" +lemma parts_mono: "G \ H \ parts(G) \ parts(H)" apply auto apply (erule parts.induct) apply (blast dest: parts.Fst parts.Snd parts.Body)+ @@ -113,7 +113,7 @@ text\Equations hold because constructors are injective.\ -lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" +lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x\A)" by auto lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" @@ -143,7 +143,7 @@ by (unfold keysFor_def, blast) text\Monotonicity\ -lemma keysFor_mono: "G \ H ==> keysFor(G) \ keysFor(H)" +lemma keysFor_mono: "G \ H \ keysFor(G) \ keysFor(H)" by (unfold keysFor_def, blast) lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H" @@ -165,7 +165,7 @@ lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}" by (unfold keysFor_def, auto) -lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H ==> invKey K \ keysFor H" +lemma Crypt_imp_invKey_keysFor: "Crypt K X \ H \ invKey K \ keysFor H" by (unfold keysFor_def, blast) @@ -192,11 +192,11 @@ apply (erule parts.induct, blast+) done -lemma parts_emptyE [elim!]: "X\ parts{} ==> P" +lemma parts_emptyE [elim!]: "X\ parts{} \ P" by simp text\WARNING: loops if H = {Y}, therefore must not be repeated!\ -lemma parts_singleton: "X\ parts H ==> \Y\H. X\ parts {Y}" +lemma parts_singleton: "X\ parts H \ \Y\H. X\ parts {Y}" by (erule parts.induct, fast+) @@ -252,7 +252,7 @@ subsubsection\Idempotence and transitivity\ -lemma parts_partsD [dest!]: "X\ parts (parts H) ==> X\ parts H" +lemma parts_partsD [dest!]: "X \ parts (parts H) \ X\ parts H" by (erule parts.induct, blast+) lemma parts_idem [simp]: "parts (parts H) = parts H" @@ -324,7 +324,7 @@ text\In any message, there is an upper bound N on its greatest nonce.\ -lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" +lemma msg_Nonce_supply: "\N. \n. N\n \ Nonce n \ parts {msg}" apply (induct_tac "msg") apply (simp_all (no_asm_simp) add: exI parts_insert2) txt\MPair case: blast works out the necessary sum itself!\ @@ -363,7 +363,7 @@ \ X \ analz H" (*<*) text\Monotonicity; Lemma 1 of Lowe's paper\ -lemma analz_mono: "G\H ==> analz(G) \ analz(H)" +lemma analz_mono: "G\H \ analz(G) \ analz(H)" apply auto apply (erule analz.induct) apply (auto dest: analz.Fst analz.Snd) @@ -435,7 +435,7 @@ text\Can only pull out Keys if they are not needed to decrypt the rest\ lemma analz_insert_Key [simp]: - "K \ keysFor (analz H) ==> + "K \ keysFor (analz H) \ analz (insert (Key K) H) = insert (Key K) (analz H)" apply (unfold keysFor_def) apply (rule analz_insert_eq_I) @@ -455,20 +455,20 @@ text\Can pull out enCrypted message if the Key is not known\ lemma analz_insert_Crypt: "Key (invKey K) \ analz H - ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" + \ analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)" apply (rule analz_insert_eq_I) apply (erule analz.induct, auto) done -lemma lemma1: "Key (invKey K) \ analz H ==> +lemma lemma1: "Key (invKey K) \ analz H \ analz (insert (Crypt K X) H) \ insert (Crypt K X) (analz (insert X H))" apply (rule subsetI) apply (erule_tac x = x in analz.induct, auto) done -lemma lemma2: "Key (invKey K) \ analz H ==> +lemma lemma2: "Key (invKey K) \ analz H \ insert (Crypt K X) (analz (insert X H)) \ analz (insert (Crypt K X) H)" apply auto @@ -477,7 +477,7 @@ done lemma analz_insert_Decrypt: - "Key (invKey K) \ analz H ==> + "Key (invKey K) \ analz H \ analz (insert (Crypt K X) H) = insert (Crypt K X) (analz (insert X H))" by (intro equalityI lemma1 lemma2) @@ -511,7 +511,7 @@ subsubsection\Idempotence and transitivity\ -lemma analz_analzD [dest!]: "X\ analz (analz H) ==> X\ analz H" +lemma analz_analzD [dest!]: "X\ analz (analz H) \ X\ analz H" by (erule analz.induct, blast+) lemma analz_idem [simp]: "analz (analz H) = analz H" @@ -531,13 +531,13 @@ by (erule analz_trans, blast) (*Cut can be proved easily by induction on - "Y: analz (insert X H) ==> X: analz H --> Y: analz H" + "Y \ analz (insert X H) \ X \ analz H \ Y \ analz H" *) text\This rewrite rule helps in the simplification of messages that involve the forwarding of unknown components (X). Without it, removing occurrences of X can be very complicated.\ -lemma analz_insert_eq: "X\ analz H ==> analz (insert X H) = analz H" +lemma analz_insert_eq: "X\ analz H \ analz (insert X H) = analz H" by (blast intro: analz_cut analz_insertI) @@ -556,7 +556,7 @@ by (intro equalityI analz_subset_cong, simp_all) lemma analz_insert_cong: - "analz H = analz H' ==> analz(insert X H) = analz(insert X H')" + "analz H = analz H' \ analz(insert X H) = analz(insert X H')" by (force simp only: insert_def intro!: analz_cong) text\If there are no pairs or encryptions then analz does nothing\ @@ -568,7 +568,7 @@ text\These two are obsolete (with a single Spy) but cost little to prove...\ lemma analz_UN_analz_lemma: - "X\ analz (\i\A. analz (H i)) ==> X\ analz (\i\A. H i)" + "X\ analz (\i\A. analz (H i)) \ X\ analz (\i\A. H i)" apply (erule analz.induct) apply (blast intro: analz_mono [THEN [2] rev_subsetD])+ done @@ -598,7 +598,7 @@ | Crypt [intro]: "\X \ synth H; Key K \ H\ \ Crypt K X \ synth H" (*<*) -lemma synth_mono: "G\H ==> synth(G) \ synth(H)" +lemma synth_mono: "G\H \ synth(G) \ synth(H)" by (auto, erule synth.induct, auto) inductive_cases Key_synth [elim!]: "Key K \ synth H" @@ -668,7 +668,7 @@ subsubsection\Idempotence and transitivity\ -lemma synth_synthD [dest!]: "X\ synth (synth H) ==> X\ synth H" +lemma synth_synthD [dest!]: "X\ synth (synth H) \ X\ synth H" by (erule synth.induct, blast+) lemma synth_idem: "synth (synth H) = synth H" @@ -697,7 +697,7 @@ by blast lemma Crypt_synth_eq [simp]: - "Key K \ H ==> (Crypt K X \ synth H) = (Crypt K X \ H)" + "Key K \ H \ (Crypt K X \ synth H) = (Crypt K X \ H)" by blast @@ -724,13 +724,13 @@ subsubsection\For reasoning about the Fake rule in traces\ -lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" +lemma parts_insert_subset_Un: "X \ G \ parts(insert X H) \ parts G \ parts H" by (rule subset_trans [OF parts_mono parts_Un_subset2], blast) text\More specifically for Fake. Very occasionally we could do with a version of the form @{term"parts{X} \ synth (analz H) \ parts H"}\ lemma Fake_parts_insert: - "X \ synth (analz H) ==> + "X \ synth (analz H) \ parts (insert X H) \ synth (analz H) \ parts H" apply (drule parts_insert_subset_Un) apply (simp (no_asm_use)) @@ -738,14 +738,14 @@ done lemma Fake_parts_insert_in_Un: - "[|Z \ parts (insert X H); X: synth (analz H)|] + "[|Z \ parts (insert X H); X \ synth (analz H)|] ==> Z \ synth (analz H) \ parts H" by (blast dest: Fake_parts_insert [THEN subsetD, dest]) text\@{term H} is sometimes @{term"Key ` KK \ spies evs"}, so can't put @{term "G=H"}.\ lemma Fake_analz_insert: - "X\ synth (analz G) ==> + "X \ synth (analz G) \ analz (insert X H) \ synth (analz G) \ analz (G \ H)" apply (rule subsetI) apply (subgoal_tac "x \ analz (synth (analz G) \ H) ") @@ -869,11 +869,11 @@ lemma Crypt_notin_image_Key [simp]: "Crypt K X \ Key ` A" by auto -lemma synth_analz_mono: "G\H ==> synth (analz(G)) \ synth (analz(H))" +lemma synth_analz_mono: "G\H \ synth (analz(G)) \ synth (analz(H))" by (iprover intro: synth_mono analz_mono) lemma Fake_analz_eq [simp]: - "X \ synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)" + "X \ synth(analz H) \ synth (analz (insert X H)) = synth (analz H)" apply (drule Fake_analz_insert[of _ _ "H"]) apply (simp add: synth_increasing[THEN Un_absorb2]) apply (drule synth_mono) @@ -885,18 +885,18 @@ text\Two generalizations of @{text analz_insert_eq}\ lemma gen_analz_insert_eq [rule_format]: - "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G" + "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) lemma synth_analz_insert_eq [rule_format]: "X \ synth (analz H) - ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)" + \ \G. H \ G \ (Key K \ analz (insert X G)) = (Key K \ analz G)" apply (erule synth.induct) apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) done lemma Fake_parts_sing: - "X \ synth (analz H) ==> parts{X} \ synth (analz H) \ parts H" + "X \ synth (analz H) \ parts{X} \ synth (analz H) \ parts H" apply (rule subset_trans) apply (erule_tac [2] Fake_parts_insert) apply (rule parts_mono, blast) diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Protocol/Public.thy --- a/src/Doc/Tutorial/Protocol/Public.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Protocol/Public.thy Thu Feb 15 13:04:36 2018 +0100 @@ -32,7 +32,7 @@ | initState_Friend: "initState (Friend i) = insert (Key (priK (Friend i))) (Key ` range pubK)" | initState_Spy: "initState Spy = - (Key`invKey`pubK`bad) Un (Key ` range pubK)" + (Key`invKey`pubK`bad) \ (Key ` range pubK)" end (*>*) @@ -77,14 +77,14 @@ (** "Image" equations that hold for injective functions **) -lemma invKey_image_eq[simp]: "(invKey x : invKey`A) = (x:A)" +lemma invKey_image_eq[simp]: "(invKey x \ invKey`A) = (x\A)" by auto (*holds because invKey is injective*) -lemma pubK_image_eq[simp]: "(pubK x : pubK`A) = (x:A)" +lemma pubK_image_eq[simp]: "(pubK x \ pubK`A) = (x\A)" by auto -lemma priK_pubK_image_eq[simp]: "(priK x ~: pubK`A)" +lemma priK_pubK_image_eq[simp]: "(priK x \ pubK`A)" by auto @@ -101,15 +101,15 @@ (*** Function "spies" ***) (*Agents see their own private keys!*) -lemma priK_in_initState[iff]: "Key (priK A) : initState A" +lemma priK_in_initState[iff]: "Key (priK A) \ initState A" by (induct A) auto (*All public keys are visible*) -lemma spies_pubK[iff]: "Key (pubK A) : spies evs" +lemma spies_pubK[iff]: "Key (pubK A) \ spies evs" by (induct evs) (simp_all add: imageI knows_Cons split: event.split) (*Spy sees private keys of bad agents!*) -lemma Spy_spies_bad[intro!]: "A: bad ==> Key (priK A) : spies evs" +lemma Spy_spies_bad[intro!]: "A \ bad \ Key (priK A) \ spies evs" by (induct evs) (simp_all add: imageI knows_Cons split: event.split) lemmas [iff] = spies_pubK [THEN analz.Inj] @@ -117,17 +117,17 @@ (*** Fresh nonces ***) -lemma Nonce_notin_initState[iff]: "Nonce N ~: parts (initState B)" +lemma Nonce_notin_initState[iff]: "Nonce N \ parts (initState B)" by (induct B) auto -lemma Nonce_notin_used_empty[simp]: "Nonce N ~: used []" +lemma Nonce_notin_used_empty[simp]: "Nonce N \ used []" by (simp add: used_Nil) (*** Supply fresh nonces for possibility theorems. ***) (*In any trace, there is an upper bound N on the greatest nonce in use.*) -lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \ used evs" +lemma Nonce_supply_lemma: "\N. \n. N\n \ Nonce n \ used evs" apply (induct_tac "evs") apply (rule_tac x = 0 in exI) apply (simp_all (no_asm_simp) add: used_Cons split: event.split) @@ -135,10 +135,10 @@ apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ done -lemma Nonce_supply1: "EX N. Nonce N \ used evs" +lemma Nonce_supply1: "\N. Nonce N \ used evs" by (rule Nonce_supply_lemma [THEN exE], blast) -lemma Nonce_supply: "Nonce (@ N. Nonce N \ used evs) \ used evs" +lemma Nonce_supply: "Nonce (SOME N. Nonce N \ used evs) \ used evs" apply (rule Nonce_supply_lemma [THEN exE]) apply (rule someI, fast) done @@ -146,10 +146,10 @@ (*** Specialized rewriting for the analz_image_... theorems ***) -lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H" +lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \ H" by blast -lemma insert_Key_image: "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C" +lemma insert_Key_image: "insert (Key K) (Key`KK \ C) = Key ` (insert K KK) \ C" by blast diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Rules/TPrimes.thy --- a/src/Doc/Tutorial/Rules/TPrimes.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Rules/TPrimes.thy Thu Feb 15 13:04:36 2018 +0100 @@ -96,7 +96,7 @@ definition is_gcd :: "[nat,nat,nat] \ bool" where (*gcd as a relation*) "is_gcd p m n == p dvd m \ p dvd n \ - (ALL d. d dvd m \ d dvd n \ d dvd p)" + (\d. d dvd m \ d dvd n \ d dvd p)" (*Function gcd yields the Greatest Common Divisor*) lemma is_gcd: "is_gcd (gcd m n) m n" diff -r e4e57da0583a -r 560fbd6bc047 src/Doc/Tutorial/Sets/Examples.thy --- a/src/Doc/Tutorial/Sets/Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/Doc/Tutorial/Sets/Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -155,7 +155,7 @@ by blast definition prime :: "nat set" where - "prime == {p. 1

m=1 | m=p)}" + "prime == {p. 1

m. m dvd p \ m=1 \ m=p)}" lemma "{p*q | p q. p\prime \ q\prime} = {z. \p q. z = p*q \ p\prime \ q\prime}" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/AbelCoset.thy Thu Feb 15 13:04:36 2018 +0100 @@ -554,7 +554,7 @@ .. lemma (in abelian_group_hom) hom_add [simp]: - "[| x : carrier G; y : carrier G |] + "[| x \ carrier G; y \ carrier G |] ==> h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" by (rule group_hom.hom_mult[OF a_group_hom, simplified ring_record_simps]) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/FiniteProduct.thy --- a/src/HOL/Algebra/FiniteProduct.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/FiniteProduct.thy Thu Feb 15 13:04:36 2018 +0100 @@ -22,7 +22,7 @@ for D :: "'a set" and f :: "'b => 'a => 'a" and e :: 'a where emptyI [intro]: "e \ D ==> ({}, e) \ foldSetD D f e" - | insertI [intro]: "[| x ~: A; f x y \ D; (A, y) \ foldSetD D f e |] ==> + | insertI [intro]: "[| x \ A; f x y \ D; (A, y) \ foldSetD D f e |] ==> (insert x A, f x y) \ foldSetD D f e" inductive_cases empty_foldSetDE [elim!]: "({}, x) \ foldSetD D f e" @@ -178,7 +178,7 @@ done lemma (in LCD) foldD_insert: - "[| finite A; x ~: A; x \ B; e \ D; A \ B |] ==> + "[| finite A; x \ A; x \ B; e \ D; A \ B |] ==> foldD D f e (insert x A) = f x (foldD D f e A)" apply (unfold foldD_def) apply (simp add: foldD_insert_aux) @@ -423,13 +423,13 @@ proof (intro finprod_insert) show "finite B" by fact next - show "x ~: B" by fact + show "x \ B" by fact next - assume "x ~: B" "!!i. i \ insert x B \ f i = g i" + assume "x \ B" "!!i. i \ insert x B \ f i = g i" "g \ insert x B \ carrier G" thus "f \ B \ carrier G" by fastforce next - assume "x ~: B" "!!i. i \ insert x B \ f i = g i" + assume "x \ B" "!!i. i \ insert x B \ f i = g i" "g \ insert x B \ carrier G" thus "f x \ carrier G" by fastforce qed @@ -491,8 +491,8 @@ (* The following two were contributed by Jeremy Avigad. *) lemma finprod_reindex: - "f : (h ` A) \ carrier G \ - inj_on h A ==> finprod G f (h ` A) = finprod G (%x. f (h x)) A" + "f \ (h ` A) \ carrier G \ + inj_on h A \ finprod G f (h ` A) = finprod G (\x. f (h x)) A" proof (induct A rule: infinite_finite_induct) case (infinite A) hence "\ finite (h ` A)" @@ -501,8 +501,8 @@ qed (auto simp add: Pi_def) lemma finprod_const: - assumes a [simp]: "a : carrier G" - shows "finprod G (%x. a) A = a [^] card A" + assumes a [simp]: "a \ carrier G" + shows "finprod G (\x. a) A = a [^] card A" proof (induct A rule: infinite_finite_induct) case (insert b A) show ?case diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/Group.thy Thu Feb 15 13:04:36 2018 +0100 @@ -435,7 +435,7 @@ "x \ carrier G \ x [^] (i + j::int) = x [^] i \ x [^] j" proof - have [simp]: "-i - j = -j - i" by simp - assume "x : carrier G" then + assume "x \ carrier G" then show ?thesis by (auto simp add: int_pow_def2 inv_solve_left inv_solve_right nat_add_distrib [symmetric] nat_pow_mult ) qed diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/Order.thy --- a/src/HOL/Algebra/Order.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/Order.thy Thu Feb 15 13:04:36 2018 +0100 @@ -379,7 +379,7 @@ proof - have "Upper L A \ carrier L" by simp moreover from above L have "s \ Upper L A" by (simp add: Upper_def) - moreover from below have "ALL x : Upper L A. s \ x" by fast + moreover from below have "\x \ Upper L A. s \ x" by fast ultimately show ?thesis by (simp add: least_def) qed @@ -439,7 +439,7 @@ proof - have "Lower L A \ carrier L" by simp moreover from below L have "i \ Lower L A" by (simp add: Lower_def) - moreover from above have "ALL x : Lower L A. x \ i" by fast + moreover from above have "\x \ Lower L A. x \ i" by fast ultimately show ?thesis by (simp add: greatest_def) qed diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/Ring.thy --- a/src/HOL/Algebra/Ring.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/Ring.thy Thu Feb 15 13:04:36 2018 +0100 @@ -74,7 +74,7 @@ and a_comm: "!!x y. [| x \ carrier R; y \ carrier R |] ==> x \ y = y \ x" and l_zero: "!!x. x \ carrier R ==> \ \ x = x" - and l_inv_ex: "!!x. x \ carrier R ==> EX y : carrier R. y \ x = \" + and l_inv_ex: "\x. x \ carrier R \ \y \ carrier R. y \ x = \" shows "abelian_group R" by (auto intro!: abelian_group.intro abelian_monoidI abelian_group_axioms.intro comm_monoidI comm_groupI diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/RingHom.thy --- a/src/HOL/Algebra/RingHom.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/RingHom.thy Thu Feb 15 13:04:36 2018 +0100 @@ -39,8 +39,8 @@ assumes "ring R" "ring S" assumes (* morphism: "h \ carrier R \ carrier S" *) hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" - and compatible_mult: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" - and compatible_add: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" + and compatible_mult: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" + and compatible_add: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" shows "ring_hom_ring R S h" proof - @@ -72,7 +72,7 @@ lemma ring_hom_ringI3: fixes R (structure) and S (structure) assumes "abelian_group_hom R S h" "ring R" "ring S" - assumes compatible_mult: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" + assumes compatible_mult: "\x y. [| x \ carrier R; y \ carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" shows "ring_hom_ring R S h" proof - diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/Sylow.thy --- a/src/HOL/Algebra/Sylow.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/Sylow.thy Thu Feb 15 13:04:36 2018 +0100 @@ -257,7 +257,7 @@ lemmas H_elem_map_eq = H_elem_map [THEN someI_ex, THEN conjunct2] lemma rcosets_H_funcset_M: - "(\C \ rcosets H. M1 #> (@g. g \ carrier G \ H #> g = C)) \ rcosets H \ M" + "(\C \ rcosets H. M1 #> (SOME g. g \ carrier G \ H #> g = C)) \ rcosets H \ M" apply (simp add: RCOSETS_def) apply (fast intro: someI2 intro!: M1_in_M in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g]) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Algebra/UnivPoly.thy Thu Feb 15 13:04:36 2018 +0100 @@ -284,7 +284,7 @@ lemma UP_l_neg_ex: assumes R: "p \ carrier P" - shows "EX q : carrier P. q \\<^bsub>P\<^esub> p = \\<^bsub>P\<^esub>" + shows "\q \ carrier P. q \\<^bsub>P\<^esub> p = \\<^bsub>P\<^esub>" proof - let ?q = "\i. \ (p i)" from R have closed: "?q \ carrier P" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Arcwise_Connected.thy --- a/src/HOL/Analysis/Arcwise_Connected.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Arcwise_Connected.thy Thu Feb 15 13:04:36 2018 +0100 @@ -21,12 +21,12 @@ where "inj B" "\n. open(B n)" and open_cov: "\S. open S \ \K. S = \(B ` K)" by (metis Setcompr_eq_image that univ_second_countable_sequence) define A where "A \ rec_nat S (\n a. if \U. U \ a \ closed U \ \ U \ U \ (B n) = {} - then @U. U \ a \ closed U \ \ U \ U \ (B n) = {} + then SOME U. U \ a \ closed U \ \ U \ U \ (B n) = {} else a)" have [simp]: "A 0 = S" by (simp add: A_def) have ASuc: "A(Suc n) = (if \U. U \ A n \ closed U \ \ U \ U \ (B n) = {} - then @U. U \ A n \ closed U \ \ U \ U \ (B n) = {} + then SOME U. U \ A n \ closed U \ \ U \ U \ (B n) = {} else A n)" for n by (auto simp: A_def) have sub: "\n. A(Suc n) \ A n" @@ -1801,7 +1801,7 @@ and peq: "\x y. \x \ T; y \ T; open_segment x y \ T = {}\ \ p x = p y" unfolding \_def by metis+ then have "T \ {}" by auto - define h where "h \ \x. p(@y. y \ T \ open_segment x y \ T = {})" + define h where "h \ \x. p(SOME y. y \ T \ open_segment x y \ T = {})" have "p y = p z" if "y \ T" "z \ T" and xyT: "open_segment x y \ T = {}" and xzT: "open_segment x z \ T = {}" for x y z proof (cases "x \ T") diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Thu Feb 15 13:04:36 2018 +0100 @@ -714,9 +714,9 @@ where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition contour_integral - where "contour_integral g f \ @i. (f has_contour_integral i) g \ ~ f contour_integrable_on g \ i=0" - -lemma not_integrable_contour_integral: "~ f contour_integrable_on g \ contour_integral g f = 0" + where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" + +lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" unfolding contour_integrable_on_def contour_integral_def by blast lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" @@ -3327,7 +3327,7 @@ definition winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ - @n. \e > 0. \p. valid_path p \ z \ path_image p \ + SOME n. \e > 0. \p. valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Continuous_Extension.thy Thu Feb 15 13:04:36 2018 +0100 @@ -309,7 +309,7 @@ "\x. x \ S \ g x = f x" proof (cases "S = {}") case True then show thesis - apply (rule_tac g="\x. @y. y \ C" in that) + apply (rule_tac g="\x. SOME y. y \ C" in that) apply (rule continuous_intros) apply (meson all_not_in_conv \C \ {}\ image_subsetI someI_ex, simp) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1944,7 +1944,7 @@ using \u + v = 1\ by auto ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \ (\x. a + x) ` S" using h1 by auto - then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto + then have "u *\<^sub>R x + v *\<^sub>R y \ S" by auto } then show ?thesis unfolding affine_def by auto qed @@ -2031,7 +2031,7 @@ have [simp]: "(\x. x - a) = plus (- a)" by (simp add: fun_eq_iff) have "affine ((\x. (-a)+x) ` S)" using affine_translation assms by auto - moreover have "0 : ((\x. (-a)+x) ` S)" + moreover have "0 \ ((\x. (-a)+x) ` S)" using assms exI[of "(\x. x\S \ -a+x = 0)" a] by auto ultimately show ?thesis using subspace_affine by auto qed @@ -2130,7 +2130,7 @@ lemma mem_cone: assumes "cone S" "x \ S" "c \ 0" - shows "c *\<^sub>R x : S" + shows "c *\<^sub>R x \ S" using assms cone_def[of S] by auto lemma cone_contains_0: @@ -3409,7 +3409,7 @@ assumes "a \ S" shows "affine_dependent (insert a S) \ dependent ((\x. -a + x) ` S)" proof - - have "((+) (- a) ` S) = {x - a| x . x : S}" by auto + have "((+) (- a) ` S) = {x - a| x . x \ S}" by auto then show ?thesis using affine_dependent_translation_eq[of "(insert a S)" "-a"] affine_dependent_imp_dependent2 assms diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1324,7 +1324,7 @@ using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def apply force+ done have le_meaT: "measure lebesgue (\\') \ measure lebesgue T" - proof (rule measure_mono_fmeasurable [OF _ _ \T : lmeasurable\]) + proof (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) show "(\\') \ sets lebesgue" using div lmeasurable_division by auto have "\\' \ \\" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Extended_Real_Limits.thy --- a/src/HOL/Analysis/Extended_Real_Limits.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Extended_Real_Limits.thy Thu Feb 15 13:04:36 2018 +0100 @@ -159,7 +159,7 @@ then obtain b where b: "Inf S - e < b" "b < Inf S" using fin ereal_between[of "Inf S" e] dense[of "Inf S - e"] by auto - then have "b: {Inf S - e <..< Inf S + e}" + then have "b \ {Inf S - e <..< Inf S + e}" using e fin ereal_between[of "Inf S" e] by auto then have "b \ S" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Polytope.thy Thu Feb 15 13:04:36 2018 +0100 @@ -3201,7 +3201,7 @@ finite \ \ (\S \ \. \n. n simplex S) \ (\F S. S \ \ \ F face_of S \ F \ \) \ - (!S S'. S \ \ \ S' \ \ + (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S \ (S \ S') face_of S')" definition triangulation where @@ -3350,7 +3350,7 @@ and convex\: "\C. C \ \ \ convex C" and closed\: "\C. C \ \ \ closed C" by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) - have in_rel_interior: "(@z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C + have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) have *: "\T. ~affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K @@ -3396,7 +3396,7 @@ by fastforce qed let ?\ = "(\C \ \. \K \ \ \ Pow (rel_frontier C). - {convex hull (insert (@z. z \ rel_interior C) K)})" + {convex hull (insert (SOME z. z \ rel_interior C) K)})" have "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ (\C \ \. \F. F \ \ \ C = \F) \ @@ -3415,9 +3415,9 @@ using S face\ that by blast moreover have "F \ \ \ ?\" if "F face_of S" "C \ \" "K \ \" and "K \ rel_frontier C" - and S: "S = convex hull insert (@z. z \ rel_interior C) K" for C K + and S: "S = convex hull insert (SOME z. z \ rel_interior C) K" for C K proof - - let ?z = "@z. z \ rel_interior C" + let ?z = "SOME z. z \ rel_interior C" have "?z \ rel_interior C" by (simp add: in_rel_interior \C \ \\) moreover @@ -3490,13 +3490,13 @@ proof - obtain C K where "C \ \" "K \ \" "K \ rel_frontier C" - and Y: "Y = convex hull insert (@z. z \ rel_interior C) K" + and Y: "Y = convex hull insert (SOME z. z \ rel_interior C) K" using XY by blast have "convex C" by (simp add: \C \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_iff) - let ?z = "(@z. z \ rel_interior C)" + let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast obtain D where "D \ \" "X \ D" @@ -3533,11 +3533,11 @@ proof - obtain C K D L where "C \ \" "K \ \" "K \ rel_frontier C" - and X: "X = convex hull insert (@z. z \ rel_interior C) K" + and X: "X = convex hull insert (SOME z. z \ rel_interior C) K" and "D \ \" "L \ \" "L \ rel_frontier D" - and Y: "Y = convex hull insert (@z. z \ rel_interior D) L" + and Y: "Y = convex hull insert (SOME z. z \ rel_interior D) L" using XY by blast - let ?z = "(@z. z \ rel_interior C)" + let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast have "convex C" @@ -3564,7 +3564,7 @@ by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_eq) have "L \ D" by (metis DiffE \D \ \\ \L \ rel_frontier D\ closed\ closure_closed rel_frontier_def subset_eq) - let ?w = "(@w. w \ rel_interior D)" + let ?w = "(SOME w. w \ rel_interior D)" have w: "?w \ rel_interior D" using \D \ \\ in_rel_interior by blast have "C \ rel_interior D = (D \ C) \ rel_interior D" @@ -3663,7 +3663,7 @@ case False then have "C \ \" by (simp add: \_def \_def aff\ less_le that) - let ?z = "@z. z \ rel_interior C" + let ?z = "SOME z. z \ rel_interior C" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast let ?F = "\K \ \ \ Pow (rel_frontier C). {convex hull (insert ?z K)}" @@ -3726,7 +3726,7 @@ next assume "L \ ?\" then obtain C K where "C \ \" - and L: "L = convex hull insert (@z. z \ rel_interior C) K" + and L: "L = convex hull insert (SOME z. z \ rel_interior C) K" and K: "K \ \" "K \ rel_frontier C" by auto then have "convex hull C = C" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Starlike.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1272,7 +1272,7 @@ fix x :: "'a::euclidean_space" fix u assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" "(\x\?D. u x *\<^sub>R x) = x" - have *: "\i\Basis. i:d \ u i = x\i" + have *: "\i\Basis. i \ d \ u i = x\i" and "(\i\Basis. i \ d \ x \ i = 0)" using as(3) unfolding substdbasis_expansion_unique[OF assms] @@ -1590,7 +1590,7 @@ case True have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] - using Min_gr_iff[of "(\) x ` d" "norm (x - y)"] \0 < card d\ \i:d\ + using Min_gr_iff[of "(\) x ` d" "norm (x - y)"] \0 < card d\ \i \ d\ by (simp add: card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] @@ -1833,7 +1833,7 @@ then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all - then have *: "x - e1 *\<^sub>R (x - a) : rel_interior S" + then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" @@ -2442,11 +2442,11 @@ subsubsection \Relative interior and closure under common operations\ -lemma rel_interior_inter_aux: "\{rel_interior S |S. S : I} \ \I" +lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y - assume "y \ \{rel_interior S |S. S : I}" + assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { @@ -2824,13 +2824,13 @@ fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto - then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 : S" + then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 \linear f\ by (simp add: linear_add_cmul) - ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S" + ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto - then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S" + then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" @@ -2861,7 +2861,7 @@ { fix z assume "z \ f -` (rel_interior S)" - then have z: "f z : rel_interior S" + then have z: "f z \ rel_interior S" by auto { fix x @@ -3115,7 +3115,7 @@ by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto - then have "y \ rel_interior {t. f t \ {}}" and "(z : rel_interior (f y))" + then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast @@ -3238,7 +3238,7 @@ have "?lhs \ ?rhs" proof fix x - assume "x : ?rhs" + assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" @@ -3676,7 +3676,7 @@ by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { - fix i assume "i:I" + fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Feb 15 13:04:36 2018 +0100 @@ -2353,7 +2353,7 @@ by (meson Int_mono closure_mono closure_subset order_refl) qed -lemma islimpt_in_closure: "(x islimpt S) = (x:closure(S-{x}))" +lemma islimpt_in_closure: "(x islimpt S) = (x\closure(S-{x}))" unfolding closure_def using islimpt_punctured by blast lemma connected_imp_connected_closure: "connected S \ connected (closure S)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/CertifiedEmail.thy --- a/src/HOL/Auth/CertifiedEmail.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/CertifiedEmail.thy Thu Feb 15 13:04:36 2018 +0100 @@ -11,7 +11,7 @@ "TTP == Server" abbreviation - RPwd :: "agent => key" where + RPwd :: "agent \ key" where "RPwd == shrK" @@ -24,7 +24,7 @@ BothAuth :: nat text\We formalize a fixed way of computing responses. Could be better.\ -definition "response" :: "agent => agent => nat => msg" where +definition "response" :: "agent \ agent \ nat \ msg" where "response S R q == Hash \Agent S, Key (shrK R), Nonce q\" @@ -129,9 +129,9 @@ lemma hr_form_lemma [rule_format]: "evs \ certified_mail - ==> hr \ synth (analz (spies evs)) --> + \ hr \ synth (analz (spies evs)) \ (\S2TTP. Notes TTP \Agent R, Agent TTP, S2TTP, pwd, hr\ - \ set evs --> + \ set evs \ (\clt q S em. hr = Hash \Number clt, Nonce q, response S R q, em\))" apply (erule certified_mail.induct) apply (synth_analz_mono_contra, simp_all, blast+) @@ -201,7 +201,7 @@ done lemma Spy_dont_know_RPwd [rule_format]: - "evs \ certified_mail ==> Key (RPwd A) \ parts(spies evs) --> A \ bad" + "evs \ certified_mail ==> Key (RPwd A) \ parts(spies evs) \ A \ bad" apply (erule certified_mail.induct, simp_all) txt\Fake\ apply (blast dest: Fake_parts_insert_in_Un) @@ -247,8 +247,8 @@ lemma analz_image_freshK [rule_format]: "evs \ certified_mail ==> - \K KK. invKey (pubEK TTP) \ KK --> - (Key K \ analz (Key`KK Un (spies evs))) = + \K KK. invKey (pubEK TTP) \ KK \ + (Key K \ analz (Key`KK \ (spies evs))) = (K \ KK | Key K \ analz (spies evs))" apply (erule certified_mail.induct) apply (drule_tac [6] A=TTP in symKey_neq_priEK) @@ -281,11 +281,11 @@ isn't inductive: message 3 case can't be proved *) lemma S2TTP_sender_lemma [rule_format]: "evs \ certified_mail ==> - Key K \ analz (spies evs) --> + Key K \ analz (spies evs) \ (\AO. Crypt (pubEK TTP) - \Agent S, Number AO, Key K, Agent R, hs\ \ used evs --> + \Agent S, Number AO, Key K, Agent R, hs\ \ used evs \ (\m ctxt q. - hs = Hash\Number ctxt, Nonce q, response S R q, Crypt K (Number m)\ & + hs = Hash\Number ctxt, Nonce q, response S R q, Crypt K (Number m)\ \ Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number ctxt, Nonce q, @@ -314,7 +314,7 @@ Key K \ analz (spies evs); evs \ certified_mail|] ==> \m ctxt q. - hs = Hash\Number ctxt, Nonce q, response S R q, Crypt K (Number m)\ & + hs = Hash\Number ctxt, Nonce q, response S R q, Crypt K (Number m)\ \ Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number ctxt, Nonce q, @@ -345,19 +345,19 @@ where @{term K} is secure.\ lemma Key_unique_lemma [rule_format]: "evs \ certified_mail ==> - Key K \ analz (spies evs) --> + Key K \ analz (spies evs) \ (\m cleartext q hs. Says S R \Agent S, Agent TTP, Crypt K (Number m), Number AO, Number cleartext, Nonce q, Crypt (pubEK TTP) \Agent S, Number AO, Key K, Agent R, hs\\ - \ set evs --> + \ set evs \ (\m' cleartext' q' hs'. Says S' R' \Agent S', Agent TTP, Crypt K (Number m'), Number AO', Number cleartext', Nonce q', Crypt (pubEK TTP) \Agent S', Number AO', Key K, Agent R', hs'\\ - \ set evs --> R' = R & S' = S & AO' = AO & hs' = hs))" + \ set evs \ R' = R \ S' = S \ AO' = AO \ hs' = hs))" apply (erule certified_mail.induct, analz_mono_contra, simp_all) prefer 2 txt\Message 1\ @@ -380,7 +380,7 @@ \ set evs; Key K \ analz (spies evs); evs \ certified_mail|] - ==> R' = R & S' = S & AO' = AO & hs' = hs" + ==> R' = R \ S' = S \ AO' = AO \ hs' = hs" by (rule Key_unique_lemma, assumption+) @@ -396,7 +396,7 @@ Key K \ analz (spies evs); evs \ certified_mail; S\Spy|] - ==> R \ bad & Gets S (Crypt (priSK TTP) S2TTP) \ set evs" + ==> R \ bad \ Gets S (Crypt (priSK TTP) S2TTP) \ set evs" apply (erule rev_mp) apply (erule ssubst) apply (erule rev_mp) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Event.thy --- a/src/HOL/Auth/Event.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Event.thy Thu Feb 15 13:04:36 2018 +0100 @@ -13,7 +13,7 @@ theory Event imports Message begin consts (*Initial states of agents -- parameter of the construction*) - initState :: "agent => msg set" + initState :: "agent \ msg set" datatype event = Says agent agent msg @@ -29,24 +29,24 @@ Server_not_bad [iff]: "Server \ bad" by (rule exI [of _ "{Spy}"], simp) -primrec knows :: "agent => event list => msg set" +primrec knows :: "agent \ event list \ msg set" where knows_Nil: "knows A [] = initState A" | knows_Cons: "knows A (ev # evs) = (if A = Spy then (case ev of - Says A' B X => insert X (knows Spy evs) - | Gets A' X => knows Spy evs - | Notes A' X => + Says A' B X \ insert X (knows Spy evs) + | Gets A' X \ knows Spy evs + | Notes A' X \ if A' \ bad then insert X (knows Spy evs) else knows Spy evs) else (case ev of - Says A' B X => + Says A' B X \ if A'=A then insert X (knows A evs) else knows A evs - | Gets A' X => + | Gets A' X \ if A'=A then insert X (knows A evs) else knows A evs - | Notes A' X => + | Notes A' X \ if A'=A then insert X (knows A evs) else knows A evs))" (* Case A=Spy on the Gets event @@ -57,31 +57,31 @@ text\The constant "spies" is retained for compatibility's sake\ abbreviation (input) - spies :: "event list => msg set" where + spies :: "event list \ msg set" where "spies == knows Spy" (*Set of items that might be visible to somebody: complement of the set of fresh items*) -primrec used :: "event list => msg set" +primrec used :: "event list \ msg set" where used_Nil: "used [] = (UN B. parts (initState B))" | used_Cons: "used (ev # evs) = (case ev of - Says A B X => parts {X} \ used evs - | Gets A X => used evs - | Notes A X => parts {X} \ used evs)" + Says A B X \ parts {X} \ used evs + | Gets A X \ used evs + | Notes A X \ parts {X} \ used evs)" \ \The case for @{term Gets} seems anomalous, but @{term Gets} always follows @{term Says} in real protocols. Seems difficult to change. See \Gets_correct\ in theory \Guard/Extensions.thy\.\ -lemma Notes_imp_used [rule_format]: "Notes A X \ set evs --> X \ used evs" +lemma Notes_imp_used [rule_format]: "Notes A X \ set evs \ X \ used evs" apply (induct_tac evs) apply (auto split: event.split) done -lemma Says_imp_used [rule_format]: "Says A B X \ set evs --> X \ used evs" +lemma Says_imp_used [rule_format]: "Says A B X \ set evs \ X \ used evs" apply (induct_tac evs) apply (auto split: event.split) done @@ -102,7 +102,7 @@ on whether @{term "A=Spy"} and whether @{term "A\bad"}\ lemma knows_Spy_Notes [simp]: "knows Spy (Notes A X # evs) = - (if A:bad then insert X (knows Spy evs) else knows Spy evs)" + (if A\bad then insert X (knows Spy evs) else knows Spy evs)" by simp lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" @@ -122,13 +122,13 @@ text\Spy sees what is sent on the traffic\ lemma Says_imp_knows_Spy [rule_format]: - "Says A B X \ set evs --> X \ knows Spy evs" + "Says A B X \ set evs \ X \ knows Spy evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done lemma Notes_imp_knows_Spy [rule_format]: - "Notes A X \ set evs --> A: bad --> X \ knows Spy evs" + "Notes A X \ set evs \ A \ bad \ X \ knows Spy evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done @@ -162,14 +162,14 @@ by (simp add: subset_insertI) text\Agents know what they say\ -lemma Says_imp_knows [rule_format]: "Says A B X \ set evs --> X \ knows A evs" +lemma Says_imp_knows [rule_format]: "Says A B X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) apply blast done text\Agents know what they note\ -lemma Notes_imp_knows [rule_format]: "Notes A X \ set evs --> X \ knows A evs" +lemma Notes_imp_knows [rule_format]: "Notes A X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) apply blast @@ -177,7 +177,7 @@ text\Agents know what they receive\ lemma Gets_imp_knows_agents [rule_format]: - "A \ Spy --> Gets A X \ set evs --> X \ knows A evs" + "A \ Spy \ Gets A X \ set evs \ X \ knows A evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) done @@ -186,8 +186,8 @@ text\What agents DIFFERENT FROM Spy know was either said, or noted, or got, or known initially\ lemma knows_imp_Says_Gets_Notes_initState [rule_format]: - "[| X \ knows A evs; A \ Spy |] ==> EX B. - Says A B X \ set evs | Gets A X \ set evs | Notes A X \ set evs | X \ initState A" + "[| X \ knows A evs; A \ Spy |] ==> \ B. + Says A B X \ set evs \ Gets A X \ set evs \ Notes A X \ set evs \ X \ initState A" apply (erule rev_mp) apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) @@ -197,8 +197,8 @@ text\What the Spy knows -- for the time being -- was either said or noted, or known initially\ lemma knows_Spy_imp_Says_Notes_initState [rule_format]: - "[| X \ knows Spy evs |] ==> EX A B. - Says A B X \ set evs | Notes A X \ set evs | X \ initState Spy" + "X \ knows Spy evs \ \A B. + Says A B X \ set evs \ Notes A X \ set evs \ X \ initState Spy" apply (erule rev_mp) apply (induct_tac "evs") apply (simp_all (no_asm_simp) split: event.split) @@ -212,7 +212,7 @@ lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] -lemma initState_into_used: "X \ parts (initState B) ==> X \ used evs" +lemma initState_into_used: "X \ parts (initState B) \ X \ used evs" apply (induct_tac "evs") apply (simp_all add: parts_insert_knows_A split: event.split, blast) done @@ -236,7 +236,7 @@ used_Nil [simp del] used_Cons [simp del] -text\For proving theorems of the form @{term "X \ analz (knows Spy evs) --> P"} +text\For proving theorems of the form @{term "X \ analz (knows Spy evs) \ P"} New events added by induction to "evs" are discarded. Provided this information isn't needed, the proof will be much shorter, since it will omit complicated reasoning about @{term analz}.\ @@ -278,7 +278,7 @@ method_setup analz_mono_contra = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\ - "for proving theorems of the form X \ analz (knows Spy evs) --> P" + "for proving theorems of the form X \ analz (knows Spy evs) \ P" subsubsection\Useful for case analysis on whether a hash is a spoof or not\ @@ -299,6 +299,6 @@ method_setup synth_analz_mono_contra = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (synth_analz_mono_contra_tac ctxt)))\ - "for proving theorems of the form X \ synth (analz (knows Spy evs)) --> P" + "for proving theorems of the form X \ synth (analz (knows Spy evs)) \ P" end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Analz.thy --- a/src/HOL/Auth/Guard/Analz.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Analz.thy Thu Feb 15 13:04:36 2018 +0100 @@ -16,37 +16,37 @@ pparts :: "msg set => msg set" for H :: "msg set" where - Inj [intro]: "[| X:H; is_MPair X |] ==> X:pparts H" -| Fst [dest]: "[| \X,Y\:pparts H; is_MPair X |] ==> X:pparts H" -| Snd [dest]: "[| \X,Y\:pparts H; is_MPair Y |] ==> Y:pparts H" + Inj [intro]: "[| X \ H; is_MPair X |] ==> X \ pparts H" +| Fst [dest]: "[| \X,Y\ \ pparts H; is_MPair X |] ==> X \ pparts H" +| Snd [dest]: "[| \X,Y\ \ pparts H; is_MPair Y |] ==> Y \ pparts H" subsection\basic facts about @{term pparts}\ -lemma pparts_is_MPair [dest]: "X:pparts H ==> is_MPair X" +lemma pparts_is_MPair [dest]: "X \ pparts H \ is_MPair X" by (erule pparts.induct, auto) -lemma Crypt_notin_pparts [iff]: "Crypt K X ~:pparts H" +lemma Crypt_notin_pparts [iff]: "Crypt K X \ pparts H" by auto -lemma Key_notin_pparts [iff]: "Key K ~:pparts H" +lemma Key_notin_pparts [iff]: "Key K \ pparts H" by auto -lemma Nonce_notin_pparts [iff]: "Nonce n ~:pparts H" +lemma Nonce_notin_pparts [iff]: "Nonce n \ pparts H" by auto -lemma Number_notin_pparts [iff]: "Number n ~:pparts H" +lemma Number_notin_pparts [iff]: "Number n \ pparts H" by auto -lemma Agent_notin_pparts [iff]: "Agent A ~:pparts H" +lemma Agent_notin_pparts [iff]: "Agent A \ pparts H" by auto lemma pparts_empty [iff]: "pparts {} = {}" by (auto, erule pparts.induct, auto) -lemma pparts_insertI [intro]: "X:pparts H ==> X:pparts (insert Y H)" +lemma pparts_insertI [intro]: "X \ pparts H \ X \ pparts (insert Y H)" by (erule pparts.induct, auto) -lemma pparts_sub: "[| X:pparts G; G<=H |] ==> X:pparts H" +lemma pparts_sub: "[| X \ pparts G; G \ H |] ==> X \ pparts H" by (erule pparts.induct, auto) lemma pparts_insert2 [iff]: "pparts (insert X (insert Y H)) @@ -78,13 +78,13 @@ lemma pparts_insert_Hash [iff]: "pparts (insert (Hash X) H) = pparts H" by (rule eq, erule pparts.induct, auto) -lemma pparts_insert: "X:pparts (insert Y H) ==> X:pparts {Y} Un pparts H" +lemma pparts_insert: "X \ pparts (insert Y H) \ X \ pparts {Y} \ pparts H" by (erule pparts.induct, blast+) -lemma insert_pparts: "X:pparts {Y} Un pparts H ==> X:pparts (insert Y H)" +lemma insert_pparts: "X \ pparts {Y} \ pparts H \ X \ pparts (insert Y H)" by (safe, erule pparts.induct, auto) -lemma pparts_Un [iff]: "pparts (G Un H) = pparts G Un pparts H" +lemma pparts_Un [iff]: "pparts (G \ H) = pparts G \ pparts H" by (rule eq, erule pparts.induct, auto dest: pparts_sub) lemma pparts_pparts [iff]: "pparts (pparts H) = pparts H" @@ -95,21 +95,21 @@ lemmas pparts_insert_substI = pparts_insert_eq [THEN ssubst] -lemma in_pparts: "Y:pparts H ==> EX X. X:H & Y:pparts {X}" +lemma in_pparts: "Y \ pparts H \ \X. X \ H \ Y \ pparts {X}" by (erule pparts.induct, auto) subsection\facts about @{term pparts} and @{term parts}\ -lemma pparts_no_Nonce [dest]: "[| X:pparts {Y}; Nonce n ~:parts {Y} |] -==> Nonce n ~:parts {X}" +lemma pparts_no_Nonce [dest]: "[| X \ pparts {Y}; Nonce n \ parts {Y} |] +==> Nonce n \ parts {X}" by (erule pparts.induct, simp_all) subsection\facts about @{term pparts} and @{term analz}\ -lemma pparts_analz: "X:pparts H ==> X:analz H" +lemma pparts_analz: "X \ pparts H \ X \ analz H" by (erule pparts.induct, auto) -lemma pparts_analz_sub: "[| X:pparts G; G<=H |] ==> X:analz H" +lemma pparts_analz_sub: "[| X \ pparts G; G \ H |] ==> X \ analz H" by (auto dest: pparts_sub pparts_analz) subsection\messages that contribute to analz\ @@ -118,23 +118,23 @@ kparts :: "msg set => msg set" for H :: "msg set" where - Inj [intro]: "[| X:H; not_MPair X |] ==> X:kparts H" -| Fst [intro]: "[| \X,Y\ \ pparts H; not_MPair X |] ==> X:kparts H" -| Snd [intro]: "[| \X,Y\ \ pparts H; not_MPair Y |] ==> Y:kparts H" + Inj [intro]: "[| X \ H; not_MPair X |] ==> X \ kparts H" +| Fst [intro]: "[| \X,Y\ \ pparts H; not_MPair X |] ==> X \ kparts H" +| Snd [intro]: "[| \X,Y\ \ pparts H; not_MPair Y |] ==> Y \ kparts H" subsection\basic facts about @{term kparts}\ -lemma kparts_not_MPair [dest]: "X:kparts H ==> not_MPair X" +lemma kparts_not_MPair [dest]: "X \ kparts H \ not_MPair X" by (erule kparts.induct, auto) lemma kparts_empty [iff]: "kparts {} = {}" by (rule eq, erule kparts.induct, auto) -lemma kparts_insertI [intro]: "X:kparts H ==> X:kparts (insert Y H)" +lemma kparts_insertI [intro]: "X \ kparts H \ X \ kparts (insert Y H)" by (erule kparts.induct, auto dest: pparts_insertI) lemma kparts_insert2 [iff]: "kparts (insert X (insert Y H)) -= kparts {X} Un kparts {Y} Un kparts H" += kparts {X} \ kparts {Y} \ kparts H" by (rule eq, (erule kparts.induct, auto)+) lemma kparts_insert_MPair [iff]: "kparts (insert \X,Y\ H) @@ -165,17 +165,17 @@ = insert (Hash X) (kparts H)" by (rule eq, erule kparts.induct, auto) -lemma kparts_insert: "X:kparts (insert X H) ==> X:kparts {X} Un kparts H" +lemma kparts_insert: "X \ kparts (insert X H) \ X \ kparts {X} \ kparts H" by (erule kparts.induct, (blast dest: pparts_insert)+) -lemma kparts_insert_fst [rule_format,dest]: "X:kparts (insert Z H) ==> -X ~:kparts H --> X:kparts {Z}" +lemma kparts_insert_fst [rule_format,dest]: "X \ kparts (insert Z H) \ +X \ kparts H \ X \ kparts {Z}" by (erule kparts.induct, (blast dest: pparts_insert)+) -lemma kparts_sub: "[| X:kparts G; G<=H |] ==> X:kparts H" +lemma kparts_sub: "[| X \ kparts G; G \ H |] ==> X \ kparts H" by (erule kparts.induct, auto dest: pparts_sub) -lemma kparts_Un [iff]: "kparts (G Un H) = kparts G Un kparts H" +lemma kparts_Un [iff]: "kparts (G \ H) = kparts G \ kparts H" by (rule eq, erule kparts.induct, auto dest: kparts_sub) lemma pparts_kparts [iff]: "pparts (kparts H) = {}" @@ -184,12 +184,12 @@ lemma kparts_kparts [iff]: "kparts (kparts H) = kparts H" by (rule eq, erule kparts.induct, auto) -lemma kparts_insert_eq: "kparts (insert X H) = kparts {X} Un kparts H" +lemma kparts_insert_eq: "kparts (insert X H) = kparts {X} \ kparts H" by (rule_tac A=H in insert_Un, rule kparts_Un) lemmas kparts_insert_substI = kparts_insert_eq [THEN ssubst] -lemma in_kparts: "Y:kparts H ==> EX X. X:H & Y:kparts {X}" +lemma in_kparts: "Y \ kparts H \ \X. X \ H \ Y \ kparts {X}" by (erule kparts.induct, auto dest: in_pparts) lemma kparts_has_no_pair [iff]: "has_no_pair (kparts H)" @@ -197,59 +197,59 @@ subsection\facts about @{term kparts} and @{term parts}\ -lemma kparts_no_Nonce [dest]: "[| X:kparts {Y}; Nonce n ~:parts {Y} |] -==> Nonce n ~:parts {X}" +lemma kparts_no_Nonce [dest]: "[| X \ kparts {Y}; Nonce n \ parts {Y} |] +==> Nonce n \ parts {X}" by (erule kparts.induct, auto) -lemma kparts_parts: "X:kparts H ==> X:parts H" +lemma kparts_parts: "X \ kparts H \ X \ parts H" by (erule kparts.induct, auto dest: pparts_analz) -lemma parts_kparts: "X:parts (kparts H) ==> X:parts H" +lemma parts_kparts: "X \ parts (kparts H) \ X \ parts H" by (erule parts.induct, auto dest: kparts_parts intro: parts.Fst parts.Snd parts.Body) -lemma Crypt_kparts_Nonce_parts [dest]: "[| Crypt K Y:kparts {Z}; -Nonce n:parts {Y} |] ==> Nonce n:parts {Z}" +lemma Crypt_kparts_Nonce_parts [dest]: "[| Crypt K Y \ kparts {Z}; +Nonce n \ parts {Y} |] ==> Nonce n \ parts {Z}" by auto subsection\facts about @{term kparts} and @{term analz}\ -lemma kparts_analz: "X:kparts H ==> X:analz H" +lemma kparts_analz: "X \ kparts H \ X \ analz H" by (erule kparts.induct, auto dest: pparts_analz) -lemma kparts_analz_sub: "[| X:kparts G; G<=H |] ==> X:analz H" +lemma kparts_analz_sub: "[| X \ kparts G; G \ H |] ==> X \ analz H" by (erule kparts.induct, auto dest: pparts_analz_sub) -lemma analz_kparts [rule_format,dest]: "X:analz H ==> -Y:kparts {X} --> Y:analz H" +lemma analz_kparts [rule_format,dest]: "X \ analz H \ +Y \ kparts {X} \ Y \ analz H" by (erule analz.induct, auto dest: kparts_analz_sub) -lemma analz_kparts_analz: "X:analz (kparts H) ==> X:analz H" +lemma analz_kparts_analz: "X \ analz (kparts H) \ X \ analz H" by (erule analz.induct, auto dest: kparts_analz) -lemma analz_kparts_insert: "X:analz (kparts (insert Z H)) ==> X:analz (kparts {Z} Un kparts H)" +lemma analz_kparts_insert: "X \ analz (kparts (insert Z H)) \ X \ analz (kparts {Z} \ kparts H)" by (rule analz_sub, auto) -lemma Nonce_kparts_synth [rule_format]: "Y:synth (analz G) -==> Nonce n:kparts {Y} --> Nonce n:analz G" +lemma Nonce_kparts_synth [rule_format]: "Y \ synth (analz G) +\ Nonce n \ kparts {Y} \ Nonce n \ analz G" by (erule synth.induct, auto) -lemma kparts_insert_synth: "[| Y:parts (insert X G); X:synth (analz G); -Nonce n:kparts {Y}; Nonce n ~:analz G |] ==> Y:parts G" +lemma kparts_insert_synth: "[| Y \ parts (insert X G); X \ synth (analz G); +Nonce n \ kparts {Y}; Nonce n \ analz G |] ==> Y \ parts G" apply (drule parts_insert_substD, clarify) apply (drule in_sub, drule_tac X=Y in parts_sub, simp) apply (auto dest: Nonce_kparts_synth) done lemma Crypt_insert_synth: - "[| Crypt K Y:parts (insert X G); X:synth (analz G); Nonce n:kparts {Y}; Nonce n ~:analz G |] - ==> Crypt K Y:parts G" + "[| Crypt K Y \ parts (insert X G); X \ synth (analz G); Nonce n \ kparts {Y}; Nonce n \ analz G |] + ==> Crypt K Y \ parts G" by (metis Fake_parts_insert_in_Un Nonce_kparts_synth UnE analz_conj_parts synth_simps(5)) subsection\analz is pparts + analz of kparts\ -lemma analz_pparts_kparts: "X:analz H ==> X:pparts H | X:analz (kparts H)" +lemma analz_pparts_kparts: "X \ analz H \ X \ pparts H \ X \ analz (kparts H)" by (erule analz.induct, auto) lemma analz_pparts_kparts_eq: "analz H = pparts H Un analz (kparts H)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Extensions.thy --- a/src/HOL/Auth/Guard/Extensions.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Extensions.thy Thu Feb 15 13:04:36 2018 +0100 @@ -11,13 +11,13 @@ subsection\Extensions to Theory \Set\\ -lemma eq: "[| !!x. x:A ==> x:B; !!x. x:B ==> x:A |] ==> A=B" +lemma eq: "[| \x. x\A \ x\B; \x. x\B \ x\A |] ==> A=B" by auto -lemma insert_Un: "P ({x} Un A) ==> P (insert x A)" +lemma insert_Un: "P ({x} \ A) \ P (insert x A)" by simp -lemma in_sub: "x:A ==> {x}<=A" +lemma in_sub: "x\A \ {x}\A" by auto @@ -51,7 +51,7 @@ subsubsection\messages that are pairs\ definition is_MPair :: "msg => bool" where -"is_MPair X == EX Y Z. X = \Y,Z\" +"is_MPair X == \Y Z. X = \Y,Z\" declare is_MPair_def [simp] @@ -86,7 +86,7 @@ declare is_MPair_def [simp del] definition has_no_pair :: "msg set => bool" where -"has_no_pair H == ALL X Y. \X,Y\ \ H" +"has_no_pair H == \X Y. \X,Y\ \ H" declare has_no_pair_def [simp] @@ -98,38 +98,38 @@ lemma wf_Crypt2 [iff]: "X ~= Crypt K X" by (induct X, auto) -lemma parts_size: "X:parts {Y} ==> X=Y | size X < size Y" +lemma parts_size: "X \ parts {Y} \ X=Y \ size X < size Y" by (erule parts.induct, auto) -lemma wf_Crypt_parts [iff]: "Crypt K X ~:parts {X}" +lemma wf_Crypt_parts [iff]: "Crypt K X \ parts {X}" by (auto dest: parts_size) subsubsection\lemmas on keysFor\ definition usekeys :: "msg set => key set" where -"usekeys G == {K. EX Y. Crypt K Y:G}" +"usekeys G \ {K. \Y. Crypt K Y \ G}" lemma finite_keysFor [intro]: "finite G ==> finite (keysFor G)" apply (simp add: keysFor_def) apply (rule finite_imageI) apply (induct G rule: finite_induct) apply auto -apply (case_tac "EX K X. x = Crypt K X", clarsimp) -apply (subgoal_tac "{Ka. EX Xa. (Ka=K & Xa=X) | Crypt Ka Xa:F} +apply (case_tac "\K X. x = Crypt K X", clarsimp) +apply (subgoal_tac "{Ka. \Xa. (Ka=K \ Xa=X) \ Crypt Ka Xa \ F} = insert K (usekeys F)", auto simp: usekeys_def) -by (subgoal_tac "{K. EX X. Crypt K X = x | Crypt K X:F} = usekeys F", +by (subgoal_tac "{K. \X. Crypt K X = x \ Crypt K X \ F} = usekeys F", auto simp: usekeys_def) subsubsection\lemmas on parts\ -lemma parts_sub: "[| X:parts G; G<=H |] ==> X:parts H" +lemma parts_sub: "[| X \ parts G; G \ H |] ==> X \ parts H" by (auto dest: parts_mono) -lemma parts_Diff [dest]: "X:parts (G - H) ==> X:parts G" +lemma parts_Diff [dest]: "X \ parts (G - H) \ X \ parts G" by (erule parts_sub, auto) -lemma parts_Diff_notin: "[| Y ~:H; Nonce n ~:parts (H - {Y}) |] -==> Nonce n ~:parts H" +lemma parts_Diff_notin: "[| Y \ H; Nonce n \ parts (H - {Y}) |] +==> Nonce n \ parts H" by simp lemmas parts_insert_substI = parts_insert [THEN ssubst] @@ -138,39 +138,39 @@ lemma finite_parts_msg [iff]: "finite (parts {X})" by (induct X, auto) -lemma finite_parts [intro]: "finite H ==> finite (parts H)" +lemma finite_parts [intro]: "finite H \ finite (parts H)" apply (erule finite_induct, simp) by (rule parts_insert_substI, simp) -lemma parts_parts: "[| X:parts {Y}; Y:parts G |] ==> X:parts G" +lemma parts_parts: "[| X \ parts {Y}; Y \ parts G |] ==> X \ parts G" by (frule parts_cut, auto) -lemma parts_parts_parts: "[| X:parts {Y}; Y:parts {Z}; Z:parts G |] ==> X:parts G" +lemma parts_parts_parts: "[| X \ parts {Y}; Y \ parts {Z}; Z \ parts G |] ==> X \ parts G" by (auto dest: parts_parts) -lemma parts_parts_Crypt: "[| Crypt K X:parts G; Nonce n:parts {X} |] -==> Nonce n:parts G" +lemma parts_parts_Crypt: "[| Crypt K X \ parts G; Nonce n \ parts {X} |] +==> Nonce n \ parts G" by (blast intro: parts.Body dest: parts_parts) subsubsection\lemmas on synth\ -lemma synth_sub: "[| X:synth G; G<=H |] ==> X:synth H" +lemma synth_sub: "[| X \ synth G; G \ H |] ==> X \ synth H" by (auto dest: synth_mono) -lemma Crypt_synth [rule_format]: "[| X:synth G; Key K ~:G |] ==> -Crypt K Y:parts {X} --> Crypt K Y:parts G" +lemma Crypt_synth [rule_format]: "[| X \ synth G; Key K \ G |] ==> +Crypt K Y \ parts {X} \ Crypt K Y \ parts G" by (erule synth.induct, auto dest: parts_sub) subsubsection\lemmas on analz\ -lemma analz_UnI1 [intro]: "X:analz G ==> X:analz (G Un H)" +lemma analz_UnI1 [intro]: "X \ analz G ==> X \ analz (G \ H)" by (subgoal_tac "G <= G Un H") (blast dest: analz_mono)+ -lemma analz_sub: "[| X:analz G; G <= H |] ==> X:analz H" +lemma analz_sub: "[| X \ analz G; G \ H |] ==> X \ analz H" by (auto dest: analz_mono) -lemma analz_Diff [dest]: "X:analz (G - H) ==> X:analz G" +lemma analz_Diff [dest]: "X \ analz (G - H) \ X \ analz G" by (erule analz.induct, auto) lemmas in_analz_subset_cong = analz_subset_cong [THEN subsetD] @@ -181,32 +181,32 @@ lemmas insert_commute_substI = insert_commute [THEN ssubst] lemma analz_insertD: - "[| Crypt K Y:H; Key (invKey K):H |] ==> analz (insert Y H) = analz H" + "[| Crypt K Y \ H; Key (invKey K) \ H |] ==> analz (insert Y H) = analz H" by (blast intro: analz.Decrypt analz_insert_eq) -lemma must_decrypt [rule_format,dest]: "[| X:analz H; has_no_pair H |] ==> -X ~:H --> (EX K Y. Crypt K Y:H & Key (invKey K):H)" +lemma must_decrypt [rule_format,dest]: "[| X \ analz H; has_no_pair H |] ==> +X \ H \ (\K Y. Crypt K Y \ H \ Key (invKey K) \ H)" by (erule analz.induct, auto) -lemma analz_needs_only_finite: "X:analz H ==> EX G. G <= H & finite G" +lemma analz_needs_only_finite: "X \ analz H \ \G. G \ H \ finite G" by (erule analz.induct, auto) -lemma notin_analz_insert: "X ~:analz (insert Y G) ==> X ~:analz G" +lemma notin_analz_insert: "X \ analz (insert Y G) \ X \ analz G" by auto subsubsection\lemmas on parts, synth and analz\ -lemma parts_invKey [rule_format,dest]:"X:parts {Y} ==> -X:analz (insert (Crypt K Y) H) --> X ~:analz H --> Key (invKey K):analz H" +lemma parts_invKey [rule_format,dest]:"X \ parts {Y} \ +X \ analz (insert (Crypt K Y) H) \ X \ analz H \ Key (invKey K) \ analz H" by (erule parts.induct, auto dest: parts.Fst parts.Snd parts.Body) -lemma in_analz: "Y:analz H ==> EX X. X:H & Y:parts {X}" +lemma in_analz: "Y \ analz H \ \X. X \ H \ Y \ parts {X}" by (erule analz.induct, auto intro: parts.Fst parts.Snd parts.Body) lemmas in_analz_subset_parts = analz_subset_parts [THEN subsetD] -lemma Crypt_synth_insert: "[| Crypt K X:parts (insert Y H); -Y:synth (analz H); Key K ~:analz H |] ==> Crypt K X:parts H" +lemma Crypt_synth_insert: "[| Crypt K X \ parts (insert Y H); +Y \ synth (analz H); Key K \ analz H |] ==> Crypt K X \ parts H" apply (drule parts_insert_substD, clarify) apply (frule in_sub) apply (frule parts_mono) @@ -222,24 +222,24 @@ | "greatest_msg (Crypt K X) = greatest_msg X" | "greatest_msg other = 0" -lemma greatest_msg_is_greatest: "Nonce n:parts {X} ==> n <= greatest_msg X" +lemma greatest_msg_is_greatest: "Nonce n \ parts {X} \ n \ greatest_msg X" by (induct X, auto) subsubsection\sets of keys\ definition keyset :: "msg set => bool" where -"keyset G == ALL X. X:G --> (EX K. X = Key K)" +"keyset G \ \X. X \ G \ (\K. X = Key K)" -lemma keyset_in [dest]: "[| keyset G; X:G |] ==> EX K. X = Key K" +lemma keyset_in [dest]: "[| keyset G; X \ G |] ==> \K. X = Key K" by (auto simp: keyset_def) lemma MPair_notin_keyset [simp]: "keyset G ==> \X,Y\ \ G" by auto -lemma Crypt_notin_keyset [simp]: "keyset G ==> Crypt K X ~:G" +lemma Crypt_notin_keyset [simp]: "keyset G \ Crypt K X \ G" by auto -lemma Nonce_notin_keyset [simp]: "keyset G ==> Nonce n ~:G" +lemma Nonce_notin_keyset [simp]: "keyset G \ Nonce n \ G" by auto lemma parts_keyset [simp]: "keyset G ==> parts G = G" @@ -256,10 +256,10 @@ lemma keyset_Diff_keysfor [simp]: "keyset H ==> keyset (H - keysfor G)" by (auto simp: keyset_def) -lemma keysfor_Crypt: "Crypt K X:parts G ==> Key (invKey K):keysfor G" +lemma keysfor_Crypt: "Crypt K X \ parts G \ Key (invKey K) \ keysfor G" by (auto simp: keysfor_def Crypt_imp_invKey_keysFor) -lemma no_key_no_Crypt: "Key K ~:keysfor G ==> Crypt (invKey K) X ~:parts G" +lemma no_key_no_Crypt: "Key K \ keysfor G \ Crypt (invKey K) X \ parts G" by (auto dest: keysfor_Crypt) lemma finite_keysfor [intro]: "finite G ==> finite (keysfor G)" @@ -273,7 +273,7 @@ apply (erule analz.induct, blast) apply (simp, blast) apply (simp, blast) -apply (case_tac "Key (invKey K):H - keysfor G", clarsimp) +apply (case_tac "Key (invKey K) \ H - keysfor G", clarsimp) apply (drule_tac X=X in no_key_no_Crypt) by (auto intro: analz_sub) @@ -286,7 +286,7 @@ subsubsection\general protocol properties\ definition is_Says :: "event => bool" where -"is_Says ev == (EX A B X. ev = Says A B X)" +"is_Says ev == (\A B X. ev = Says A B X)" lemma is_Says_Says [iff]: "is_Says (Says A B X)" by (simp add: is_Says_def) @@ -294,36 +294,36 @@ (* one could also require that Gets occurs after Says but this is sufficient for our purpose *) definition Gets_correct :: "event list set => bool" where -"Gets_correct p == ALL evs B X. evs:p --> Gets B X:set evs ---> (EX A. Says A B X:set evs)" +"Gets_correct p == \evs B X. evs \ p \ Gets B X \ set evs +\ (\A. Says A B X \ set evs)" -lemma Gets_correct_Says: "[| Gets_correct p; Gets B X # evs:p |] -==> EX A. Says A B X:set evs" +lemma Gets_correct_Says: "[| Gets_correct p; Gets B X # evs \ p |] +==> \A. Says A B X \ set evs" apply (simp add: Gets_correct_def) by (drule_tac x="Gets B X # evs" in spec, auto) definition one_step :: "event list set => bool" where -"one_step p == ALL evs ev. ev#evs:p --> evs:p" +"one_step p == \evs ev. ev#evs \ p \ evs \ p" -lemma one_step_Cons [dest]: "[| one_step p; ev#evs:p |] ==> evs:p" +lemma one_step_Cons [dest]: "[| one_step p; ev#evs \ p |] ==> evs \ p" by (unfold one_step_def, blast) -lemma one_step_app: "[| evs@evs':p; one_step p; []:p |] ==> evs':p" +lemma one_step_app: "[| evs@evs' \ p; one_step p; [] \ p |] ==> evs' \ p" by (induct evs, auto) -lemma trunc: "[| evs @ evs':p; one_step p |] ==> evs':p" +lemma trunc: "[| evs @ evs' \ p; one_step p |] ==> evs' \ p" by (induct evs, auto) definition has_only_Says :: "event list set => bool" where -"has_only_Says p == ALL evs ev. evs:p --> ev:set evs ---> (EX A B X. ev = Says A B X)" +"has_only_Says p \ \evs ev. evs \ p \ ev \ set evs +\ (\A B X. ev = Says A B X)" -lemma has_only_SaysD: "[| ev:set evs; evs:p; has_only_Says p |] -==> EX A B X. ev = Says A B X" +lemma has_only_SaysD: "[| ev \ set evs; evs \ p; has_only_Says p |] +==> \A B X. ev = Says A B X" by (unfold has_only_Says_def, blast) -lemma in_has_only_Says [dest]: "[| has_only_Says p; evs:p; ev:set evs |] -==> EX A B X. ev = Says A B X" +lemma in_has_only_Says [dest]: "[| has_only_Says p; evs \ p; ev \ set evs |] +==> \A B X. ev = Says A B X" by (auto simp: has_only_Says_def) lemma has_only_Says_imp_Gets_correct [simp]: "has_only_Says p @@ -332,11 +332,11 @@ subsubsection\lemma on knows\ -lemma Says_imp_spies2: "Says A B \X,Y\ \ set evs ==> Y \ parts (spies evs)" +lemma Says_imp_spies2: "Says A B \X,Y\ \ set evs \ Y \ parts (spies evs)" by (drule Says_imp_spies, drule parts.Inj, drule parts.Snd, simp) -lemma Says_not_parts: "[| Says A B X:set evs; Y ~:parts (spies evs) |] -==> Y ~:parts {X}" +lemma Says_not_parts: "[| Says A B X \ set evs; Y \ parts (spies evs) |] +==> Y \ parts {X}" by (auto dest: Says_imp_spies parts_parts) subsubsection\knows without initState\ @@ -349,7 +349,7 @@ case ev of Says A' B X => insert X (knows' A evs) | Gets A' X => knows' A evs - | Notes A' X => if A':bad then insert X (knows' A evs) else knows' A evs + | Notes A' X => if A' \ bad then insert X (knows' A evs) else knows' A evs ) else ( case ev of Says A' B X => if A=A' then insert X (knows' A evs) else knows' A evs @@ -390,8 +390,8 @@ lemmas knows_Cons_substI = knows_Cons [THEN ssubst] lemmas knows_Cons_substD = knows_Cons [THEN sym, THEN ssubst] -lemma knows'_sub_spies': "[| evs:p; has_only_Says p; one_step p |] -==> knows' A evs <= spies' evs" +lemma knows'_sub_spies': "[| evs \ p; has_only_Says p; one_step p |] +==> knows' A evs \ spies' evs" by (induct evs, auto split: event.splits) subsubsection\knows' is finite\ @@ -404,7 +404,7 @@ lemma knows_sub_Cons: "knows A evs <= knows A (ev#evs)" by(cases A, induct evs, auto simp: knows.simps split:event.split) -lemma knows_ConsI: "X:knows A evs ==> X:knows A (ev#evs)" +lemma knows_ConsI: "X \ knows A evs \ X \ knows A (ev#evs)" by (auto dest: knows_sub_Cons [THEN subsetD]) lemma knows_sub_app: "knows A evs <= knows A (evs @ evs')" @@ -424,7 +424,7 @@ Says A' B X => insert X (knows_max' A evs) | Gets A' X => knows_max' A evs | Notes A' X => - if A':bad then insert X (knows_max' A evs) else knows_max' A evs + if A' \ bad then insert X (knows_max' A evs) else knows_max' A evs ) else ( case ev of Says A' B X => @@ -466,22 +466,22 @@ lemma finite_knows_max' [iff]: "finite (knows_max' A evs)" by (induct evs, auto split: event.split) -lemma knows_max'_sub_spies': "[| evs:p; has_only_Says p; one_step p |] -==> knows_max' A evs <= spies' evs" +lemma knows_max'_sub_spies': "[| evs \ p; has_only_Says p; one_step p |] +==> knows_max' A evs \ spies' evs" by (induct evs, auto split: event.splits) -lemma knows_max'_in_spies' [dest]: "[| evs:p; X:knows_max' A evs; -has_only_Says p; one_step p |] ==> X:spies' evs" +lemma knows_max'_in_spies' [dest]: "[| evs \ p; X \ knows_max' A evs; +has_only_Says p; one_step p |] ==> X \ spies' evs" by (rule knows_max'_sub_spies' [THEN subsetD], auto) lemma knows_max'_app: "knows_max' A (evs @ evs') = knows_max' A evs Un knows_max' A evs'" by (induct evs, auto split: event.splits) -lemma Says_to_knows_max': "Says A B X:set evs ==> X:knows_max' B evs" +lemma Says_to_knows_max': "Says A B X \ set evs \ X \ knows_max' B evs" by (simp add: in_set_conv_decomp, clarify, simp add: knows_max'_app) -lemma Says_from_knows_max': "Says A B X:set evs ==> X:knows_max' A evs" +lemma Says_from_knows_max': "Says A B X \ set evs \ X \ knows_max' A evs" by (simp add: in_set_conv_decomp, clarify, simp add: knows_max'_app) subsubsection\used without initState\ @@ -501,10 +501,10 @@ lemma used_decomp: "used evs = init Un used' evs" by (induct evs, auto simp: init_def split: event.split) -lemma used'_sub_app: "used' evs <= used' (evs@evs')" +lemma used'_sub_app: "used' evs \ used' (evs@evs')" by (induct evs, auto split: event.split) -lemma used'_parts [rule_format]: "X:used' evs ==> Y:parts {X} --> Y:used' evs" +lemma used'_parts [rule_format]: "X \ used' evs \ Y \ parts {X} \ Y \ used' evs" apply (induct evs, simp) apply (rename_tac a b) apply (case_tac a, simp_all) @@ -516,35 +516,35 @@ lemma used_sub_Cons: "used evs <= used (ev#evs)" by (induct evs, (induct ev, auto)+) -lemma used_ConsI: "X:used evs ==> X:used (ev#evs)" +lemma used_ConsI: "X \ used evs \ X \ used (ev#evs)" by (auto dest: used_sub_Cons [THEN subsetD]) -lemma notin_used_ConsD: "X ~:used (ev#evs) ==> X ~:used evs" +lemma notin_used_ConsD: "X \ used (ev#evs) \ X \ used evs" by (auto dest: used_sub_Cons [THEN subsetD]) -lemma used_appD [dest]: "X:used (evs @ evs') ==> X:used evs | X:used evs'" +lemma used_appD [dest]: "X \ used (evs @ evs') \ X \ used evs \ X \ used evs'" by (induct evs, auto, rename_tac a b, case_tac a, auto) -lemma used_ConsD: "X:used (ev#evs) ==> X:used [ev] | X:used evs" +lemma used_ConsD: "X \ used (ev#evs) \ X \ used [ev] \ X \ used evs" by (case_tac ev, auto) lemma used_sub_app: "used evs <= used (evs@evs')" by (auto simp: used_decomp dest: used'_sub_app [THEN subsetD]) -lemma used_appIL: "X:used evs ==> X:used (evs' @ evs)" +lemma used_appIL: "X \ used evs \ X \ used (evs' @ evs)" by (induct evs', auto intro: used_ConsI) -lemma used_appIR: "X:used evs ==> X:used (evs @ evs')" +lemma used_appIR: "X \ used evs \ X \ used (evs @ evs')" by (erule used_sub_app [THEN subsetD]) -lemma used_parts: "[| X:parts {Y}; Y:used evs |] ==> X:used evs" +lemma used_parts: "[| X \ parts {Y}; Y \ used evs |] ==> X \ used evs" apply (auto simp: used_decomp dest: used'_parts) by (auto simp: init_def used_Nil dest: parts_trans) -lemma parts_Says_used: "[| Says A B X:set evs; Y:parts {X} |] ==> Y:used evs" +lemma parts_Says_used: "[| Says A B X \ set evs; Y \ parts {X} |] ==> Y \ used evs" by (induct evs, simp_all, safe, auto intro: used_ConsI) -lemma parts_used_app: "X:parts {Y} ==> X:used (evs @ Says A B Y # evs')" +lemma parts_used_app: "X \ parts {Y} \ X \ used (evs @ Says A B Y # evs')" apply (drule_tac evs="[Says A B Y]" in used_parts, simp, blast) apply (drule_tac evs'=evs' in used_appIR) apply (drule_tac evs'=evs in used_appIL) @@ -552,67 +552,67 @@ subsubsection\lemmas on used and knows\ -lemma initState_used: "X:parts (initState A) ==> X:used evs" +lemma initState_used: "X \ parts (initState A) \ X \ used evs" by (induct evs, auto simp: used.simps split: event.split) -lemma Says_imp_used: "Says A B X:set evs ==> parts {X} <= used evs" +lemma Says_imp_used: "Says A B X \ set evs \ parts {X} \ used evs" by (induct evs, auto intro: used_ConsI) -lemma not_used_not_spied: "X ~:used evs ==> X ~:parts (spies evs)" +lemma not_used_not_spied: "X \ used evs \ X \ parts (spies evs)" by (induct evs, auto simp: used_Nil) -lemma not_used_not_parts: "[| Y ~:used evs; Says A B X:set evs |] -==> Y ~:parts {X}" +lemma not_used_not_parts: "[| Y \ used evs; Says A B X \ set evs |] +==> Y \ parts {X}" by (induct evs, auto intro: used_ConsI) -lemma not_used_parts_false: "[| X ~:used evs; Y:parts (spies evs) |] -==> X ~:parts {Y}" +lemma not_used_parts_false: "[| X \ used evs; Y \ parts (spies evs) |] +==> X \ parts {Y}" by (auto dest: parts_parts) -lemma known_used [rule_format]: "[| evs:p; Gets_correct p; one_step p |] -==> X:parts (knows A evs) --> X:used evs" +lemma known_used [rule_format]: "[| evs \ p; Gets_correct p; one_step p |] +==> X \ parts (knows A evs) \ X \ used evs" apply (case_tac "A=Spy", blast) apply (induct evs) apply (simp add: used.simps, blast) apply (rename_tac a evs) apply (frule_tac ev=a and evs=evs in one_step_Cons, simp, clarify) -apply (drule_tac P="%G. X:parts G" in knows_Cons_substD, safe) +apply (drule_tac P="\G. X \ parts G" in knows_Cons_substD, safe) apply (erule initState_used) apply (case_tac a, auto) apply (rename_tac msg) apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says) by (auto dest: Says_imp_used intro: used_ConsI) -lemma known_max_used [rule_format]: "[| evs:p; Gets_correct p; one_step p |] -==> X:parts (knows_max A evs) --> X:used evs" +lemma known_max_used [rule_format]: "[| evs \ p; Gets_correct p; one_step p |] +==> X \ parts (knows_max A evs) \ X \ used evs" apply (case_tac "A=Spy") apply force apply (induct evs) apply (simp add: knows_max_def used.simps, blast) apply (rename_tac a evs) apply (frule_tac ev=a and evs=evs in one_step_Cons, simp, clarify) -apply (drule_tac P="%G. X:parts G" in knows_max_Cons_substD, safe) +apply (drule_tac P="\G. X \ parts G" in knows_max_Cons_substD, safe) apply (case_tac a, auto) apply (rename_tac msg) apply (drule_tac B=A and X=msg and evs=evs in Gets_correct_Says) by (auto simp: knows_max'_Cons dest: Says_imp_used intro: used_ConsI) -lemma not_used_not_known: "[| evs:p; X ~:used evs; -Gets_correct p; one_step p |] ==> X ~:parts (knows A evs)" +lemma not_used_not_known: "[| evs \ p; X \ used evs; +Gets_correct p; one_step p |] ==> X \ parts (knows A evs)" by (case_tac "A=Spy", auto dest: not_used_not_spied known_used) -lemma not_used_not_known_max: "[| evs:p; X ~:used evs; -Gets_correct p; one_step p |] ==> X ~:parts (knows_max A evs)" +lemma not_used_not_known_max: "[| evs \ p; X \ used evs; +Gets_correct p; one_step p |] ==> X \ parts (knows_max A evs)" by (case_tac "A=Spy", auto dest: not_used_not_spied known_max_used) subsubsection\a nonce or key in a message cannot equal a fresh nonce or key\ -lemma Nonce_neq [dest]: "[| Nonce n' ~:used evs; -Says A B X:set evs; Nonce n:parts {X} |] ==> n ~= n'" +lemma Nonce_neq [dest]: "[| Nonce n' \ used evs; +Says A B X \ set evs; Nonce n \ parts {X} |] ==> n \ n'" by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub) -lemma Key_neq [dest]: "[| Key n' ~:used evs; -Says A B X:set evs; Key n:parts {X} |] ==> n ~= n'" +lemma Key_neq [dest]: "[| Key n' \ used evs; +Says A B X \ set evs; Key n \ parts {X} |] ==> n ~= n'" by (drule not_used_not_spied, auto dest: Says_imp_knows_Spy parts_sub) subsubsection\message of an event\ @@ -623,7 +623,7 @@ | "msg (Gets A X) = X" | "msg (Notes A X) = X" -lemma used_sub_parts_used: "X:used (ev # evs) ==> X:parts {msg ev} Un used evs" +lemma used_sub_parts_used: "X \ used (ev # evs) ==> X \ parts {msg ev} \ used evs" by (induct ev, auto) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard.thy --- a/src/HOL/Auth/Guard/Guard.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard.thy Thu Feb 15 13:04:36 2018 +0100 @@ -13,82 +13,82 @@ ******************************************************************************) inductive_set - guard :: "nat => key set => msg set" + guard :: "nat \ key set \ msg set" for n :: nat and Ks :: "key set" where - No_Nonce [intro]: "Nonce n ~:parts {X} ==> X:guard n Ks" -| Guard_Nonce [intro]: "invKey K:Ks ==> Crypt K X:guard n Ks" -| Crypt [intro]: "X:guard n Ks ==> Crypt K X:guard n Ks" -| Pair [intro]: "[| X:guard n Ks; Y:guard n Ks |] ==> \X,Y\ \ guard n Ks" + No_Nonce [intro]: "Nonce n \ parts {X} \ X \ guard n Ks" +| Guard_Nonce [intro]: "invKey K \ Ks \ Crypt K X \ guard n Ks" +| Crypt [intro]: "X \ guard n Ks \ Crypt K X \ guard n Ks" +| Pair [intro]: "[| X \ guard n Ks; Y \ guard n Ks |] ==> \X,Y\ \ guard n Ks" subsection\basic facts about @{term guard}\ -lemma Key_is_guard [iff]: "Key K:guard n Ks" +lemma Key_is_guard [iff]: "Key K \ guard n Ks" by auto -lemma Agent_is_guard [iff]: "Agent A:guard n Ks" +lemma Agent_is_guard [iff]: "Agent A \ guard n Ks" by auto -lemma Number_is_guard [iff]: "Number r:guard n Ks" +lemma Number_is_guard [iff]: "Number r \ guard n Ks" by auto -lemma Nonce_notin_guard: "X:guard n Ks ==> X ~= Nonce n" +lemma Nonce_notin_guard: "X \ guard n Ks \ X \ Nonce n" by (erule guard.induct, auto) -lemma Nonce_notin_guard_iff [iff]: "Nonce n ~:guard n Ks" +lemma Nonce_notin_guard_iff [iff]: "Nonce n \ guard n Ks" by (auto dest: Nonce_notin_guard) -lemma guard_has_Crypt [rule_format]: "X:guard n Ks ==> Nonce n:parts {X} ---> (EX K Y. Crypt K Y:kparts {X} & Nonce n:parts {Y})" +lemma guard_has_Crypt [rule_format]: "X \ guard n Ks ==> Nonce n \ parts {X} +\ (\K Y. Crypt K Y \ kparts {X} \ Nonce n \ parts {Y})" by (erule guard.induct, auto) -lemma Nonce_notin_kparts_msg: "X:guard n Ks ==> Nonce n ~:kparts {X}" +lemma Nonce_notin_kparts_msg: "X \ guard n Ks \ Nonce n \ kparts {X}" by (erule guard.induct, auto) -lemma Nonce_in_kparts_imp_no_guard: "Nonce n:kparts H -==> EX X. X:H & X ~:guard n Ks" +lemma Nonce_in_kparts_imp_no_guard: "Nonce n \ kparts H +\ \X. X \ H \ X \ guard n Ks" apply (drule in_kparts, clarify) apply (rule_tac x=X in exI, clarify) by (auto dest: Nonce_notin_kparts_msg) -lemma guard_kparts [rule_format]: "X:guard n Ks ==> -Y:kparts {X} --> Y:guard n Ks" +lemma guard_kparts [rule_format]: "X \ guard n Ks \ +Y \ kparts {X} \ Y \ guard n Ks" by (erule guard.induct, auto) -lemma guard_Crypt: "[| Crypt K Y:guard n Ks; K ~:invKey`Ks |] ==> Y:guard n Ks" - by (ind_cases "Crypt K Y:guard n Ks") (auto intro!: image_eqI) +lemma guard_Crypt: "[| Crypt K Y \ guard n Ks; K \ invKey`Ks |] ==> Y \ guard n Ks" + by (ind_cases "Crypt K Y \ guard n Ks") (auto intro!: image_eqI) lemma guard_MPair [iff]: "(\X,Y\ \ guard n Ks) = (X \ guard n Ks \ Y \ guard n Ks)" by (auto, (ind_cases "\X,Y\ \ guard n Ks", auto)+) -lemma guard_not_guard [rule_format]: "X:guard n Ks ==> -Crypt K Y:kparts {X} --> Nonce n:kparts {Y} --> Y ~:guard n Ks" +lemma guard_not_guard [rule_format]: "X \ guard n Ks \ +Crypt K Y \ kparts {X} \ Nonce n \ kparts {Y} \ Y \ guard n Ks" by (erule guard.induct, auto dest: guard_kparts) -lemma guard_extand: "[| X:guard n Ks; Ks <= Ks' |] ==> X:guard n Ks'" +lemma guard_extand: "[| X \ guard n Ks; Ks \ Ks' |] ==> X \ guard n Ks'" by (erule guard.induct, auto) subsection\guarded sets\ -definition Guard :: "nat => key set => msg set => bool" where -"Guard n Ks H == ALL X. X:H --> X:guard n Ks" +definition Guard :: "nat \ key set \ msg set \ bool" where +"Guard n Ks H \ \X. X \ H \ X \ guard n Ks" subsection\basic facts about @{term Guard}\ lemma Guard_empty [iff]: "Guard n Ks {}" by (simp add: Guard_def) -lemma notin_parts_Guard [intro]: "Nonce n ~:parts G ==> Guard n Ks G" +lemma notin_parts_Guard [intro]: "Nonce n \ parts G \ Guard n Ks G" apply (unfold Guard_def, clarify) -apply (subgoal_tac "Nonce n ~:parts {X}") +apply (subgoal_tac "Nonce n \ parts {X}") by (auto dest: parts_sub) -lemma Nonce_notin_kparts [simplified]: "Guard n Ks H ==> Nonce n ~:kparts H" +lemma Nonce_notin_kparts [simplified]: "Guard n Ks H \ Nonce n \ kparts H" by (auto simp: Guard_def dest: in_kparts Nonce_notin_kparts_msg) -lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n:analz H |] ==> -EX K Y. Crypt K Y:kparts H & Key (invKey K):kparts H" -apply (drule_tac P="%G. Nonce n:G" in analz_pparts_kparts_substD, simp) +lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n \ analz H |] ==> +\K Y. Crypt K Y \ kparts H \ Key (invKey K) \ kparts H" +apply (drule_tac P="\G. Nonce n \ G" in analz_pparts_kparts_substD, simp) by (drule must_decrypt, auto dest: Nonce_notin_kparts) lemma Guard_kparts [intro]: "Guard n Ks H ==> Guard n Ks (kparts H)" @@ -98,7 +98,7 @@ by (auto simp: Guard_def) lemma Guard_insert [iff]: "Guard n Ks (insert X H) -= (Guard n Ks H & X:guard n Ks)" += (Guard n Ks H \ X \ guard n Ks)" by (auto simp: Guard_def) lemma Guard_Un [iff]: "Guard n Ks (G Un H) = (Guard n Ks G & Guard n Ks H)" @@ -107,51 +107,51 @@ lemma Guard_synth [intro]: "Guard n Ks G ==> Guard n Ks (synth G)" by (auto simp: Guard_def, erule synth.induct, auto) -lemma Guard_analz [intro]: "[| Guard n Ks G; ALL K. K:Ks --> Key K ~:analz G |] +lemma Guard_analz [intro]: "[| Guard n Ks G; \K. K \ Ks \ Key K \ analz G |] ==> Guard n Ks (analz G)" apply (auto simp: Guard_def) apply (erule analz.induct, auto) -by (ind_cases "Crypt K Xa:guard n Ks" for K Xa, auto) +by (ind_cases "Crypt K Xa \ guard n Ks" for K Xa, auto) -lemma in_Guard [dest]: "[| X:G; Guard n Ks G |] ==> X:guard n Ks" +lemma in_Guard [dest]: "[| X \ G; Guard n Ks G |] ==> X \ guard n Ks" by (auto simp: Guard_def) -lemma in_synth_Guard: "[| X:synth G; Guard n Ks G |] ==> X:guard n Ks" +lemma in_synth_Guard: "[| X \ synth G; Guard n Ks G |] ==> X \ guard n Ks" by (drule Guard_synth, auto) -lemma in_analz_Guard: "[| X:analz G; Guard n Ks G; -ALL K. K:Ks --> Key K ~:analz G |] ==> X:guard n Ks" +lemma in_analz_Guard: "[| X \ analz G; Guard n Ks G; +\K. K \ Ks \ Key K \ analz G |] ==> X \ guard n Ks" by (drule Guard_analz, auto) lemma Guard_keyset [simp]: "keyset G ==> Guard n Ks G" by (auto simp: Guard_def) -lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G Un H)" +lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G \ H)" by auto -lemma in_Guard_kparts: "[| X:G; Guard n Ks G; Y:kparts {X} |] ==> Y:guard n Ks" +lemma in_Guard_kparts: "[| X \ G; Guard n Ks G; Y \ kparts {X} |] ==> Y \ guard n Ks" by blast -lemma in_Guard_kparts_neq: "[| X:G; Guard n Ks G; Nonce n':kparts {X} |] -==> n ~= n'" +lemma in_Guard_kparts_neq: "[| X \ G; Guard n Ks G; Nonce n' \ kparts {X} |] +==> n \ n'" by (blast dest: in_Guard_kparts) -lemma in_Guard_kparts_Crypt: "[| X:G; Guard n Ks G; is_MPair X; -Crypt K Y:kparts {X}; Nonce n:kparts {Y} |] ==> invKey K:Ks" +lemma in_Guard_kparts_Crypt: "[| X \ G; Guard n Ks G; is_MPair X; +Crypt K Y \ kparts {X}; Nonce n \ kparts {Y} |] ==> invKey K \ Ks" apply (drule in_Guard, simp) apply (frule guard_not_guard, simp+) apply (drule guard_kparts, simp) -by (ind_cases "Crypt K Y:guard n Ks", auto) +by (ind_cases "Crypt K Y \ guard n Ks", auto) -lemma Guard_extand: "[| Guard n Ks G; Ks <= Ks' |] ==> Guard n Ks' G" +lemma Guard_extand: "[| Guard n Ks G; Ks \ Ks' |] ==> Guard n Ks' G" by (auto simp: Guard_def dest: guard_extand) -lemma guard_invKey [rule_format]: "[| X:guard n Ks; Nonce n:kparts {Y} |] ==> -Crypt K Y:kparts {X} --> invKey K:Ks" +lemma guard_invKey [rule_format]: "[| X \ guard n Ks; Nonce n \ kparts {Y} |] ==> +Crypt K Y \ kparts {X} \ invKey K \ Ks" by (erule guard.induct, auto) -lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y:guard n Ks; -Nonce n:kparts {Y} |] ==> invKey K:Ks" +lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y \ guard n Ks; +Nonce n \ kparts {Y} |] ==> invKey K \ Ks" by (auto dest: guard_invKey) subsection\set obtained by decrypting a message\ @@ -160,14 +160,14 @@ decrypt :: "msg set => key => msg => msg set" where "decrypt H K Y == insert Y (H - {Crypt K Y})" -lemma analz_decrypt: "[| Crypt K Y:H; Key (invKey K):H; Nonce n:analz H |] -==> Nonce n:analz (decrypt H K Y)" -apply (drule_tac P="%H. Nonce n:analz H" in ssubst [OF insert_Diff]) +lemma analz_decrypt: "[| Crypt K Y \ H; Key (invKey K) \ H; Nonce n \ analz H |] +==> Nonce n \ analz (decrypt H K Y)" +apply (drule_tac P="\H. Nonce n \ analz H" in ssubst [OF insert_Diff]) apply assumption apply (simp only: analz_Crypt_if, simp) done -lemma parts_decrypt: "[| Crypt K Y:H; X:parts (decrypt H K Y) |] ==> X:parts H" +lemma parts_decrypt: "[| Crypt K Y \ H; X \ parts (decrypt H K Y) |] ==> X \ parts H" by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body) subsection\number of Crypt's in a message\ @@ -180,7 +180,7 @@ subsection\basic facts about @{term crypt_nb}\ -lemma non_empty_crypt_msg: "Crypt K Y:parts {X} ==> crypt_nb X \ 0" +lemma non_empty_crypt_msg: "Crypt K Y \ parts {X} \ crypt_nb X \ 0" by (induct X, simp_all, safe, simp_all) subsection\number of Crypt's in a message list\ @@ -206,16 +206,16 @@ apply simp done -lemma parts_cnb: "Z:parts (set l) ==> +lemma parts_cnb: "Z \ parts (set l) \ cnb l = (cnb l - crypt_nb Z) + crypt_nb Z" by (erule parts.induct, auto simp: in_set_conv_decomp) -lemma non_empty_crypt: "Crypt K Y:parts (set l) ==> cnb l \ 0" +lemma non_empty_crypt: "Crypt K Y \ parts (set l) \ cnb l \ 0" by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD) subsection\list of kparts\ -lemma kparts_msg_set: "EX l. kparts {X} = set l & cnb l = crypt_nb X" +lemma kparts_msg_set: "\l. kparts {X} = set l \ cnb l = crypt_nb X" apply (induct X, simp_all) apply (rename_tac agent, rule_tac x="[Agent agent]" in exI, simp) apply (rename_tac nat, rule_tac x="[Number nat]" in exI, simp) @@ -225,11 +225,11 @@ apply (clarify, rule_tac x="l@la" in exI, simp) by (clarify, rename_tac nat X y, rule_tac x="[Crypt nat X]" in exI, simp) -lemma kparts_set: "EX l'. kparts (set l) = set l' & cnb l' = cnb l" +lemma kparts_set: "\l'. kparts (set l) = set l' \ cnb l' = cnb l" apply (induct l) apply (rule_tac x="[]" in exI, simp, clarsimp) apply (rename_tac a b l') -apply (subgoal_tac "EX l''. kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify) +apply (subgoal_tac "\l''. kparts {a} = set l'' \ cnb l'' = crypt_nb a", clarify) apply (rule_tac x="l''@l'" in exI, simp) apply (rule kparts_insert_substI, simp) by (rule kparts_msg_set) @@ -249,28 +249,28 @@ subsection\if the analyse of a finite guarded set gives n then it must also gives one of the keys of Ks\ -lemma Guard_invKey_by_list [rule_format]: "ALL l. cnb l = p ---> Guard n Ks (set l) --> Nonce n:analz (set l) ---> (EX K. K:Ks & Key K:analz (set l))" +lemma Guard_invKey_by_list [rule_format]: "\l. cnb l = p +\ Guard n Ks (set l) \ Nonce n \ analz (set l) +\ (\K. K \ Ks \ Key K \ analz (set l))" apply (induct p) (* case p=0 *) apply (clarify, drule Guard_must_decrypt, simp, clarify) apply (drule kparts_parts, drule non_empty_crypt, simp) (* case p>0 *) apply (clarify, frule Guard_must_decrypt, simp, clarify) -apply (drule_tac P="%G. Nonce n:G" in analz_pparts_kparts_substD, simp) +apply (drule_tac P="\G. Nonce n \ G" in analz_pparts_kparts_substD, simp) apply (frule analz_decrypt, simp_all) -apply (subgoal_tac "EX l'. kparts (set l) = set l' & cnb l' = cnb l", clarsimp) +apply (subgoal_tac "\l'. kparts (set l) = set l' \ cnb l' = cnb l", clarsimp) apply (drule_tac G="insert Y (set l' - {Crypt K Y})" and H="set (decrypt' l' K Y)" in analz_sub, rule decrypt_minus) apply (rule_tac analz_pparts_kparts_substI, simp) -apply (case_tac "K:invKey`Ks") +apply (case_tac "K \ invKey`Ks") (* K:invKey`Ks *) apply (clarsimp, blast) (* K ~:invKey`Ks *) apply (subgoal_tac "Guard n Ks (set (decrypt' l' K Y))") apply (drule_tac x="decrypt' l' K Y" in spec, simp) -apply (subgoal_tac "Crypt K Y:parts (set l)") +apply (subgoal_tac "Crypt K Y \ parts (set l)") apply (drule parts_cnb, rotate_tac -1, simp) apply (clarify, drule_tac X="Key Ka" and H="insert Y (set l')" in analz_sub) apply (rule insert_mono, rule set_remove) @@ -286,21 +286,21 @@ apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast) by (rule kparts_set) -lemma Guard_invKey_finite: "[| Nonce n:analz G; Guard n Ks G; finite G |] -==> EX K. K:Ks & Key K:analz G" +lemma Guard_invKey_finite: "[| Nonce n \ analz G; Guard n Ks G; finite G |] +==> \K. K \ Ks \ Key K \ analz G" apply (drule finite_list, clarify) by (rule Guard_invKey_by_list, auto) -lemma Guard_invKey: "[| Nonce n:analz G; Guard n Ks G |] -==> EX K. K:Ks & Key K:analz G" +lemma Guard_invKey: "[| Nonce n \ analz G; Guard n Ks G |] +==> \K. K \ Ks \ Key K \ analz G" by (auto dest: analz_needs_only_finite Guard_invKey_finite) subsection\if the analyse of a finite guarded set and a (possibly infinite) set of keys gives n then it must also gives Ks\ -lemma Guard_invKey_keyset: "[| Nonce n:analz (G Un H); Guard n Ks G; finite G; -keyset H |] ==> EX K. K:Ks & Key K:analz (G Un H)" -apply (frule_tac P="%G. Nonce n:G" and G=G in analz_keyset_substD, simp_all) +lemma Guard_invKey_keyset: "[| Nonce n \ analz (G \ H); Guard n Ks G; finite G; +keyset H |] ==> \K. K \ Ks \ Key K \ analz (G \ H)" +apply (frule_tac P="\G. Nonce n \ G" and G=G in analz_keyset_substD, simp_all) apply (drule_tac G="G Un (H Int keysfor G)" in Guard_invKey_finite) by (auto simp: Guard_def intro: analz_sub) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/GuardK.thy --- a/src/HOL/Auth/Guard/GuardK.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/GuardK.thy Thu Feb 15 13:04:36 2018 +0100 @@ -23,149 +23,149 @@ guardK :: "nat => key set => msg set" for n :: nat and Ks :: "key set" where - No_Key [intro]: "Key n ~:parts {X} ==> X:guardK n Ks" -| Guard_Key [intro]: "invKey K:Ks ==> Crypt K X:guardK n Ks" -| Crypt [intro]: "X:guardK n Ks ==> Crypt K X:guardK n Ks" -| Pair [intro]: "[| X:guardK n Ks; Y:guardK n Ks |] ==> \X,Y\:guardK n Ks" + No_Key [intro]: "Key n \ parts {X} \ X \ guardK n Ks" +| Guard_Key [intro]: "invKey K \ Ks ==> Crypt K X \ guardK n Ks" +| Crypt [intro]: "X \ guardK n Ks \ Crypt K X \ guardK n Ks" +| Pair [intro]: "[| X \ guardK n Ks; Y \ guardK n Ks |] ==> \X,Y\ \ guardK n Ks" subsection\basic facts about @{term guardK}\ -lemma Nonce_is_guardK [iff]: "Nonce p:guardK n Ks" +lemma Nonce_is_guardK [iff]: "Nonce p \ guardK n Ks" by auto -lemma Agent_is_guardK [iff]: "Agent A:guardK n Ks" +lemma Agent_is_guardK [iff]: "Agent A \ guardK n Ks" by auto -lemma Number_is_guardK [iff]: "Number r:guardK n Ks" +lemma Number_is_guardK [iff]: "Number r \ guardK n Ks" by auto -lemma Key_notin_guardK: "X:guardK n Ks ==> X ~= Key n" +lemma Key_notin_guardK: "X \ guardK n Ks \ X \ Key n" by (erule guardK.induct, auto) -lemma Key_notin_guardK_iff [iff]: "Key n ~:guardK n Ks" +lemma Key_notin_guardK_iff [iff]: "Key n \ guardK n Ks" by (auto dest: Key_notin_guardK) -lemma guardK_has_Crypt [rule_format]: "X:guardK n Ks ==> Key n:parts {X} ---> (EX K Y. Crypt K Y:kparts {X} & Key n:parts {Y})" +lemma guardK_has_Crypt [rule_format]: "X \ guardK n Ks \ Key n \ parts {X} +\ (\K Y. Crypt K Y \ kparts {X} \ Key n \ parts {Y})" by (erule guardK.induct, auto) -lemma Key_notin_kparts_msg: "X:guardK n Ks ==> Key n ~:kparts {X}" +lemma Key_notin_kparts_msg: "X \ guardK n Ks \ Key n \ kparts {X}" by (erule guardK.induct, auto dest: kparts_parts) -lemma Key_in_kparts_imp_no_guardK: "Key n:kparts H -==> EX X. X:H & X ~:guardK n Ks" +lemma Key_in_kparts_imp_no_guardK: "Key n \ kparts H +\ \X. X \ H \ X \ guardK n Ks" apply (drule in_kparts, clarify) apply (rule_tac x=X in exI, clarify) by (auto dest: Key_notin_kparts_msg) -lemma guardK_kparts [rule_format]: "X:guardK n Ks ==> -Y:kparts {X} --> Y:guardK n Ks" +lemma guardK_kparts [rule_format]: "X \ guardK n Ks \ +Y \ kparts {X} \ Y \ guardK n Ks" by (erule guardK.induct, auto dest: kparts_parts parts_sub) -lemma guardK_Crypt: "[| Crypt K Y:guardK n Ks; K ~:invKey`Ks |] ==> Y:guardK n Ks" - by (ind_cases "Crypt K Y:guardK n Ks") (auto intro!: image_eqI) +lemma guardK_Crypt: "[| Crypt K Y \ guardK n Ks; K \ invKey`Ks |] ==> Y \ guardK n Ks" + by (ind_cases "Crypt K Y \ guardK n Ks") (auto intro!: image_eqI) -lemma guardK_MPair [iff]: "(\X,Y\:guardK n Ks) -= (X:guardK n Ks & Y:guardK n Ks)" -by (auto, (ind_cases "\X,Y\:guardK n Ks", auto)+) +lemma guardK_MPair [iff]: "(\X,Y\ \ guardK n Ks) += (X \ guardK n Ks \ Y \ guardK n Ks)" +by (auto, (ind_cases "\X,Y\ \ guardK n Ks", auto)+) -lemma guardK_not_guardK [rule_format]: "X:guardK n Ks ==> -Crypt K Y:kparts {X} --> Key n:kparts {Y} --> Y ~:guardK n Ks" +lemma guardK_not_guardK [rule_format]: "X \guardK n Ks \ +Crypt K Y \ kparts {X} \ Key n \ kparts {Y} \ Y \ guardK n Ks" by (erule guardK.induct, auto dest: guardK_kparts) -lemma guardK_extand: "[| X:guardK n Ks; Ks <= Ks'; -[| K:Ks'; K ~:Ks |] ==> Key K ~:parts {X} |] ==> X:guardK n Ks'" +lemma guardK_extand: "[| X \ guardK n Ks; Ks \ Ks'; +[| K \ Ks'; K \ Ks |] ==> Key K \ parts {X} |] ==> X \ guardK n Ks'" by (erule guardK.induct, auto) subsection\guarded sets\ -definition GuardK :: "nat => key set => msg set => bool" where -"GuardK n Ks H == ALL X. X:H --> X:guardK n Ks" +definition GuardK :: "nat \ key set \ msg set \ bool" where +"GuardK n Ks H \ \X. X \ H \ X \ guardK n Ks" subsection\basic facts about @{term GuardK}\ lemma GuardK_empty [iff]: "GuardK n Ks {}" by (simp add: GuardK_def) -lemma Key_notin_kparts [simplified]: "GuardK n Ks H ==> Key n ~:kparts H" +lemma Key_notin_kparts [simplified]: "GuardK n Ks H \ Key n \ kparts H" by (auto simp: GuardK_def dest: in_kparts Key_notin_kparts_msg) -lemma GuardK_must_decrypt: "[| GuardK n Ks H; Key n:analz H |] ==> -EX K Y. Crypt K Y:kparts H & Key (invKey K):kparts H" -apply (drule_tac P="%G. Key n:G" in analz_pparts_kparts_substD, simp) +lemma GuardK_must_decrypt: "[| GuardK n Ks H; Key n \ analz H |] ==> +\K Y. Crypt K Y \ kparts H \ Key (invKey K) \ kparts H" +apply (drule_tac P="\G. Key n \ G" in analz_pparts_kparts_substD, simp) by (drule must_decrypt, auto dest: Key_notin_kparts) -lemma GuardK_kparts [intro]: "GuardK n Ks H ==> GuardK n Ks (kparts H)" +lemma GuardK_kparts [intro]: "GuardK n Ks H \ GuardK n Ks (kparts H)" by (auto simp: GuardK_def dest: in_kparts guardK_kparts) -lemma GuardK_mono: "[| GuardK n Ks H; G <= H |] ==> GuardK n Ks G" +lemma GuardK_mono: "[| GuardK n Ks H; G \ H |] ==> GuardK n Ks G" by (auto simp: GuardK_def) lemma GuardK_insert [iff]: "GuardK n Ks (insert X H) -= (GuardK n Ks H & X:guardK n Ks)" += (GuardK n Ks H \ X \ guardK n Ks)" by (auto simp: GuardK_def) lemma GuardK_Un [iff]: "GuardK n Ks (G Un H) = (GuardK n Ks G & GuardK n Ks H)" by (auto simp: GuardK_def) -lemma GuardK_synth [intro]: "GuardK n Ks G ==> GuardK n Ks (synth G)" +lemma GuardK_synth [intro]: "GuardK n Ks G \ GuardK n Ks (synth G)" by (auto simp: GuardK_def, erule synth.induct, auto) -lemma GuardK_analz [intro]: "[| GuardK n Ks G; ALL K. K:Ks --> Key K ~:analz G |] +lemma GuardK_analz [intro]: "[| GuardK n Ks G; \K. K \ Ks \ Key K \ analz G |] ==> GuardK n Ks (analz G)" apply (auto simp: GuardK_def) apply (erule analz.induct, auto) -by (ind_cases "Crypt K Xa:guardK n Ks" for K Xa, auto) +by (ind_cases "Crypt K Xa \ guardK n Ks" for K Xa, auto) -lemma in_GuardK [dest]: "[| X:G; GuardK n Ks G |] ==> X:guardK n Ks" +lemma in_GuardK [dest]: "[| X \ G; GuardK n Ks G |] ==> X \ guardK n Ks" by (auto simp: GuardK_def) -lemma in_synth_GuardK: "[| X:synth G; GuardK n Ks G |] ==> X:guardK n Ks" +lemma in_synth_GuardK: "[| X \ synth G; GuardK n Ks G |] ==> X \ guardK n Ks" by (drule GuardK_synth, auto) -lemma in_analz_GuardK: "[| X:analz G; GuardK n Ks G; -ALL K. K:Ks --> Key K ~:analz G |] ==> X:guardK n Ks" +lemma in_analz_GuardK: "[| X \ analz G; GuardK n Ks G; +\K. K \ Ks \ Key K \ analz G |] ==> X \ guardK n Ks" by (drule GuardK_analz, auto) -lemma GuardK_keyset [simp]: "[| keyset G; Key n ~:G |] ==> GuardK n Ks G" +lemma GuardK_keyset [simp]: "[| keyset G; Key n \ G |] ==> GuardK n Ks G" by (simp only: GuardK_def, clarify, drule keyset_in, auto) -lemma GuardK_Un_keyset: "[| GuardK n Ks G; keyset H; Key n ~:H |] +lemma GuardK_Un_keyset: "[| GuardK n Ks G; keyset H; Key n \ H |] ==> GuardK n Ks (G Un H)" by auto -lemma in_GuardK_kparts: "[| X:G; GuardK n Ks G; Y:kparts {X} |] ==> Y:guardK n Ks" +lemma in_GuardK_kparts: "[| X \ G; GuardK n Ks G; Y \ kparts {X} |] ==> Y \ guardK n Ks" by blast -lemma in_GuardK_kparts_neq: "[| X:G; GuardK n Ks G; Key n':kparts {X} |] -==> n ~= n'" +lemma in_GuardK_kparts_neq: "[| X \ G; GuardK n Ks G; Key n' \ kparts {X} |] +==> n \ n'" by (blast dest: in_GuardK_kparts) -lemma in_GuardK_kparts_Crypt: "[| X:G; GuardK n Ks G; is_MPair X; -Crypt K Y:kparts {X}; Key n:kparts {Y} |] ==> invKey K:Ks" +lemma in_GuardK_kparts_Crypt: "[| X \ G; GuardK n Ks G; is_MPair X; +Crypt K Y \ kparts {X}; Key n \ kparts {Y} |] ==> invKey K \ Ks" apply (drule in_GuardK, simp) apply (frule guardK_not_guardK, simp+) apply (drule guardK_kparts, simp) -by (ind_cases "Crypt K Y:guardK n Ks", auto) +by (ind_cases "Crypt K Y \ guardK n Ks", auto) -lemma GuardK_extand: "[| GuardK n Ks G; Ks <= Ks'; -[| K:Ks'; K ~:Ks |] ==> Key K ~:parts G |] ==> GuardK n Ks' G" +lemma GuardK_extand: "[| GuardK n Ks G; Ks \ Ks'; +[| K \ Ks'; K \ Ks |] ==> Key K \ parts G |] ==> GuardK n Ks' G" by (auto simp: GuardK_def dest: guardK_extand parts_sub) subsection\set obtained by decrypting a message\ abbreviation (input) - decrypt :: "msg set => key => msg => msg set" where - "decrypt H K Y == insert Y (H - {Crypt K Y})" + decrypt :: "msg set \ key \ msg \ msg set" where + "decrypt H K Y \ insert Y (H - {Crypt K Y})" -lemma analz_decrypt: "[| Crypt K Y:H; Key (invKey K):H; Key n:analz H |] -==> Key n:analz (decrypt H K Y)" -apply (drule_tac P="%H. Key n:analz H" in ssubst [OF insert_Diff]) +lemma analz_decrypt: "[| Crypt K Y \ H; Key (invKey K) \ H; Key n \ analz H |] +==> Key n \ analz (decrypt H K Y)" +apply (drule_tac P="\H. Key n \ analz H" in ssubst [OF insert_Diff]) apply assumption apply (simp only: analz_Crypt_if, simp) done -lemma parts_decrypt: "[| Crypt K Y:H; X:parts (decrypt H K Y) |] ==> X:parts H" +lemma parts_decrypt: "[| Crypt K Y \ H; X \ parts (decrypt H K Y) |] ==> X \ parts H" by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body) subsection\number of Crypt's in a message\ @@ -177,7 +177,7 @@ subsection\basic facts about @{term crypt_nb}\ -lemma non_empty_crypt_msg: "Crypt K Y:parts {X} ==> crypt_nb X \ 0" +lemma non_empty_crypt_msg: "Crypt K Y \ parts {X} \ crypt_nb X \ 0" by (induct X, simp_all, safe, simp_all) subsection\number of Crypt's in a message list\ @@ -200,16 +200,16 @@ apply (induct l, auto) by (erule_tac l=l and x=x in mem_cnb_minus_substI, simp) -lemma parts_cnb: "Z:parts (set l) ==> +lemma parts_cnb: "Z \ parts (set l) \ cnb l = (cnb l - crypt_nb Z) + crypt_nb Z" by (erule parts.induct, auto simp: in_set_conv_decomp) -lemma non_empty_crypt: "Crypt K Y:parts (set l) ==> cnb l \ 0" +lemma non_empty_crypt: "Crypt K Y \ parts (set l) \ cnb l \ 0" by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD) subsection\list of kparts\ -lemma kparts_msg_set: "EX l. kparts {X} = set l & cnb l = crypt_nb X" +lemma kparts_msg_set: "\l. kparts {X} = set l \ cnb l = crypt_nb X" apply (induct X, simp_all) apply (rename_tac agent, rule_tac x="[Agent agent]" in exI, simp) apply (rename_tac nat, rule_tac x="[Number nat]" in exI, simp) @@ -219,11 +219,11 @@ apply (clarify, rule_tac x="l@la" in exI, simp) by (clarify, rename_tac nat X y, rule_tac x="[Crypt nat X]" in exI, simp) -lemma kparts_set: "EX l'. kparts (set l) = set l' & cnb l' = cnb l" +lemma kparts_set: "\l'. kparts (set l) = set l' & cnb l' = cnb l" apply (induct l) apply (rule_tac x="[]" in exI, simp, clarsimp) apply (rename_tac a b l') -apply (subgoal_tac "EX l''. kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify) +apply (subgoal_tac "\l''. kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify) apply (rule_tac x="l''@l'" in exI, simp) apply (rule kparts_insert_substI, simp) by (rule kparts_msg_set) @@ -243,28 +243,28 @@ text\if the analysis of a finite guarded set gives n then it must also give one of the keys of Ks\ -lemma GuardK_invKey_by_list [rule_format]: "ALL l. cnb l = p ---> GuardK n Ks (set l) --> Key n:analz (set l) ---> (EX K. K:Ks & Key K:analz (set l))" +lemma GuardK_invKey_by_list [rule_format]: "\l. cnb l = p +\ GuardK n Ks (set l) \ Key n \ analz (set l) +\ (\K. K \ Ks \ Key K \ analz (set l))" apply (induct p) (* case p=0 *) apply (clarify, drule GuardK_must_decrypt, simp, clarify) apply (drule kparts_parts, drule non_empty_crypt, simp) (* case p>0 *) apply (clarify, frule GuardK_must_decrypt, simp, clarify) -apply (drule_tac P="%G. Key n:G" in analz_pparts_kparts_substD, simp) +apply (drule_tac P="\G. Key n \ G" in analz_pparts_kparts_substD, simp) apply (frule analz_decrypt, simp_all) -apply (subgoal_tac "EX l'. kparts (set l) = set l' & cnb l' = cnb l", clarsimp) +apply (subgoal_tac "\l'. kparts (set l) = set l' \ cnb l' = cnb l", clarsimp) apply (drule_tac G="insert Y (set l' - {Crypt K Y})" and H="set (decrypt' l' K Y)" in analz_sub, rule decrypt_minus) apply (rule_tac analz_pparts_kparts_substI, simp) -apply (case_tac "K:invKey`Ks") +apply (case_tac "K \ invKey`Ks") (* K:invKey`Ks *) apply (clarsimp, blast) (* K ~:invKey`Ks *) apply (subgoal_tac "GuardK n Ks (set (decrypt' l' K Y))") apply (drule_tac x="decrypt' l' K Y" in spec, simp) -apply (subgoal_tac "Crypt K Y:parts (set l)") +apply (subgoal_tac "Crypt K Y \ parts (set l)") apply (drule parts_cnb, rotate_tac -1, simp) apply (clarify, drule_tac X="Key Ka" and H="insert Y (set l')" in analz_sub) apply (rule insert_mono, rule set_remove) @@ -280,21 +280,21 @@ apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast) by (rule kparts_set) -lemma GuardK_invKey_finite: "[| Key n:analz G; GuardK n Ks G; finite G |] -==> EX K. K:Ks & Key K:analz G" +lemma GuardK_invKey_finite: "[| Key n \ analz G; GuardK n Ks G; finite G |] +==> \K. K \ Ks \ Key K \ analz G" apply (drule finite_list, clarify) by (rule GuardK_invKey_by_list, auto) -lemma GuardK_invKey: "[| Key n:analz G; GuardK n Ks G |] -==> EX K. K:Ks & Key K:analz G" +lemma GuardK_invKey: "[| Key n \ analz G; GuardK n Ks G |] +==> \K. K \ Ks \ Key K \ analz G" by (auto dest: analz_needs_only_finite GuardK_invKey_finite) text\if the analyse of a finite guarded set and a (possibly infinite) set of keys gives n then it must also gives Ks\ -lemma GuardK_invKey_keyset: "[| Key n:analz (G Un H); GuardK n Ks G; finite G; -keyset H; Key n ~:H |] ==> EX K. K:Ks & Key K:analz (G Un H)" -apply (frule_tac P="%G. Key n:G" and G=G in analz_keyset_substD, simp_all) +lemma GuardK_invKey_keyset: "[| Key n \ analz (G \ H); GuardK n Ks G; finite G; +keyset H; Key n \ H |] ==> \K. K \ Ks \ Key K \ analz (G \ H)" +apply (frule_tac P="\G. Key n \ G" and G=G in analz_keyset_substD, simp_all) apply (drule_tac G="G Un (H Int keysfor G)" in GuardK_invKey_finite) apply (auto simp: GuardK_def intro: analz_sub) by (drule keyset_in, auto) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard_NS_Public.thy --- a/src/HOL/Auth/Guard/Guard_NS_Public.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy Thu Feb 15 13:04:36 2018 +0100 @@ -37,17 +37,17 @@ inductive_set nsp :: "event list set" where - Nil: "[]:nsp" + Nil: "[] \ nsp" -| Fake: "[| evs:nsp; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs : nsp" +| Fake: "[| evs \ nsp; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ nsp" -| NS1: "[| evs1:nsp; Nonce NA ~:used evs1 |] ==> ns1 A B NA # evs1 : nsp" +| NS1: "[| evs1 \ nsp; Nonce NA \ used evs1 |] ==> ns1 A B NA # evs1 \ nsp" -| NS2: "[| evs2:nsp; Nonce NB ~:used evs2; ns1' A' A B NA:set evs2 |] ==> - ns2 B A NA NB # evs2:nsp" +| NS2: "[| evs2 \ nsp; Nonce NB \ used evs2; ns1' A' A B NA \ set evs2 |] ==> + ns2 B A NA NB # evs2 \ nsp" -| NS3: "!!A B B' NA NB evs3. [| evs3:nsp; ns1 A B NA:set evs3; ns2' B' B A NA NB:set evs3 |] ==> - ns3 A B NB # evs3:nsp" +| NS3: "\A B B' NA NB evs3. [| evs3 \ nsp; ns1 A B NA \ set evs3; ns2' B' B A NA NB \ set evs3 |] ==> + ns3 A B NB # evs3 \ nsp" subsection\declarations for tactics\ @@ -57,17 +57,17 @@ subsection\general properties of nsp\ -lemma nsp_has_no_Gets: "evs:nsp ==> ALL A X. Gets A X ~:set evs" +lemma nsp_has_no_Gets: "evs \ nsp \ \A X. Gets A X \ set evs" by (erule nsp.induct, auto) lemma nsp_is_Gets_correct [iff]: "Gets_correct nsp" by (auto simp: Gets_correct_def dest: nsp_has_no_Gets) lemma nsp_is_one_step [iff]: "one_step nsp" -by (unfold one_step_def, clarify, ind_cases "ev#evs:nsp" for ev evs, auto) +by (unfold one_step_def, clarify, ind_cases "ev#evs \ nsp" for ev evs, auto) -lemma nsp_has_only_Says' [rule_format]: "evs:nsp ==> -ev:set evs --> (EX A B X. ev=Says A B X)" +lemma nsp_has_only_Says' [rule_format]: "evs \ nsp \ +ev \ set evs \ (\A B X. ev=Says A B X)" by (erule nsp.induct, auto) lemma nsp_has_only_Says [iff]: "has_only_Says nsp" @@ -79,37 +79,37 @@ subsection\nonce are used only once\ -lemma NA_is_uniq [rule_format]: "evs:nsp ==> -Crypt (pubK B) \Nonce NA, Agent A\:parts (spies evs) ---> Crypt (pubK B') \Nonce NA, Agent A'\:parts (spies evs) ---> Nonce NA ~:analz (spies evs) --> A=A' & B=B'" +lemma NA_is_uniq [rule_format]: "evs \ nsp \ +Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs) +\ Crypt (pubK B') \Nonce NA, Agent A'\ \ parts (spies evs) +\ Nonce NA \ analz (spies evs) \ A=A' \ B=B'" apply (erule nsp.induct, simp_all) by (blast intro: analz_insertI)+ -lemma no_Nonce_NS1_NS2 [rule_format]: "evs:nsp ==> -Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\:parts (spies evs) ---> Crypt (pubK B) \Nonce NA, Agent A\:parts (spies evs) ---> Nonce NA:analz (spies evs)" +lemma no_Nonce_NS1_NS2 [rule_format]: "evs \ nsp \ +Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\ \ parts (spies evs) +\ Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs) +\ Nonce NA \ analz (spies evs)" apply (erule nsp.induct, simp_all) by (blast intro: analz_insertI)+ lemma no_Nonce_NS1_NS2' [rule_format]: -"[| Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\:parts (spies evs); -Crypt (pubK B) \Nonce NA, Agent A\:parts (spies evs); evs:nsp |] -==> Nonce NA:analz (spies evs)" +"[| Crypt (pubK B') \Nonce NA', Nonce NA, Agent A'\ \ parts (spies evs); +Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs); evs \ nsp |] +==> Nonce NA \ analz (spies evs)" by (rule no_Nonce_NS1_NS2, auto) -lemma NB_is_uniq [rule_format]: "evs:nsp ==> -Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\:parts (spies evs) ---> Crypt (pubK A') \Nonce NA', Nonce NB, Agent B'\:parts (spies evs) ---> Nonce NB ~:analz (spies evs) --> A=A' & B=B' & NA=NA'" +lemma NB_is_uniq [rule_format]: "evs \ nsp \ +Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\ \ parts (spies evs) +\ Crypt (pubK A') \Nonce NA', Nonce NB, Agent B'\ \ parts (spies evs) +\ Nonce NB \ analz (spies evs) \ A=A' \ B=B' \ NA=NA'" apply (erule nsp.induct, simp_all) by (blast intro: analz_insertI)+ subsection\guardedness of NA\ -lemma ns1_imp_Guard [rule_format]: "[| evs:nsp; A ~:bad; B ~:bad |] ==> -ns1 A B NA:set evs --> Guard NA {priK A,priK B} (spies evs)" +lemma ns1_imp_Guard [rule_format]: "[| evs \ nsp; A \ bad; B \ bad |] ==> +ns1 A B NA \ set evs \ Guard NA {priK A,priK B} (spies evs)" apply (erule nsp.induct) (* Nil *) apply simp_all @@ -135,8 +135,8 @@ subsection\guardedness of NB\ -lemma ns2_imp_Guard [rule_format]: "[| evs:nsp; A ~:bad; B ~:bad |] ==> -ns2 B A NA NB:set evs --> Guard NB {priK A,priK B} (spies evs)" +lemma ns2_imp_Guard [rule_format]: "[| evs \ nsp; A \ bad; B \ bad |] ==> +ns2 B A NA NB \ set evs \ Guard NB {priK A,priK B} (spies evs)" apply (erule nsp.induct) (* Nil *) apply simp_all @@ -165,15 +165,15 @@ subsection\Agents' Authentication\ -lemma B_trusts_NS1: "[| evs:nsp; A ~:bad; B ~:bad |] ==> -Crypt (pubK B) \Nonce NA, Agent A\:parts (spies evs) ---> Nonce NA ~:analz (spies evs) --> ns1 A B NA:set evs" +lemma B_trusts_NS1: "[| evs \ nsp; A \ bad; B \ bad |] ==> +Crypt (pubK B) \Nonce NA, Agent A\ \ parts (spies evs) +\ Nonce NA \ analz (spies evs) \ ns1 A B NA \ set evs" apply (erule nsp.induct, simp_all) by (blast intro: analz_insertI)+ -lemma A_trusts_NS2: "[| evs:nsp; A ~:bad; B ~:bad |] ==> ns1 A B NA:set evs ---> Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\:parts (spies evs) ---> ns2 B A NA NB:set evs" +lemma A_trusts_NS2: "[| evs \ nsp; A \ bad; B \ bad |] ==> ns1 A B NA \ set evs +\ Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\ \ parts (spies evs) +\ ns2 B A NA NB \ set evs" apply (erule nsp.induct, simp_all, safe) apply (frule_tac B=B in ns1_imp_Guard, simp+) apply (drule Guard_Nonce_analz, simp+, blast) @@ -182,8 +182,8 @@ apply (frule_tac B=B in ns1_imp_Guard, simp+) by (drule Guard_Nonce_analz, simp+, blast+) -lemma B_trusts_NS3: "[| evs:nsp; A ~:bad; B ~:bad |] ==> ns2 B A NA NB:set evs ---> Crypt (pubK B) (Nonce NB):parts (spies evs) --> ns3 A B NB:set evs" +lemma B_trusts_NS3: "[| evs \ nsp; A \ bad; B \ bad |] ==> ns2 B A NA NB \ set evs +\ Crypt (pubK B) (Nonce NB) \ parts (spies evs) \ ns3 A B NB \ set evs" apply (erule nsp.induct, simp_all, safe) apply (frule_tac B=B in ns2_imp_Guard, simp+) apply (drule Guard_Nonce_analz, simp+, blast) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard_OtwayRees.thy --- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy Thu Feb 15 13:04:36 2018 +0100 @@ -59,20 +59,20 @@ inductive_set or :: "event list set" where - Nil: "[]:or" + Nil: "[] \ or" -| Fake: "[| evs:or; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs:or" +| Fake: "[| evs \ or; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ or" -| OR1: "[| evs1:or; Nonce NA ~:used evs1 |] ==> or1 A B NA # evs1:or" +| OR1: "[| evs1 \ or; Nonce NA \ used evs1 |] ==> or1 A B NA # evs1 \ or" -| OR2: "[| evs2:or; or1' A' A B NA X:set evs2; Nonce NB ~:used evs2 |] - ==> or2 A B NA NB X # evs2:or" +| OR2: "[| evs2 \ or; or1' A' A B NA X \ set evs2; Nonce NB \ used evs2 |] + ==> or2 A B NA NB X # evs2 \ or" -| OR3: "[| evs3:or; or2' B' A B NA NB:set evs3; Key K ~:used evs3 |] - ==> or3 A B NA NB K # evs3:or" +| OR3: "[| evs3 \ or; or2' B' A B NA NB \ set evs3; Key K \ used evs3 |] + ==> or3 A B NA NB K # evs3 \ or" -| OR4: "[| evs4:or; or2 A B NA NB X:set evs4; or3' S Y A B NA NB K:set evs4 |] - ==> or4 A B NA X # evs4:or" +| OR4: "[| evs4 \ or; or2 A B NA NB X \ set evs4; or3' S Y A B NA NB K \ set evs4 |] + ==> or4 A B NA X # evs4 \ or" subsection\declarations for tactics\ @@ -82,17 +82,17 @@ subsection\general properties of or\ -lemma or_has_no_Gets: "evs:or ==> ALL A X. Gets A X ~:set evs" +lemma or_has_no_Gets: "evs \ or \ \A X. Gets A X \ set evs" by (erule or.induct, auto) lemma or_is_Gets_correct [iff]: "Gets_correct or" by (auto simp: Gets_correct_def dest: or_has_no_Gets) lemma or_is_one_step [iff]: "one_step or" -by (unfold one_step_def, clarify, ind_cases "ev#evs:or" for ev evs, auto) +by (unfold one_step_def, clarify, ind_cases "ev#evs \ or" for ev evs, auto) -lemma or_has_only_Says' [rule_format]: "evs:or ==> -ev:set evs --> (EX A B X. ev=Says A B X)" +lemma or_has_only_Says' [rule_format]: "evs \ or \ +ev \ set evs \ (\A B X. ev=Says A B X)" by (erule or.induct, auto) lemma or_has_only_Says [iff]: "has_only_Says or" @@ -100,16 +100,16 @@ subsection\or is regular\ -lemma or1'_parts_spies [dest]: "or1' A' A B NA X:set evs -==> X:parts (spies evs)" +lemma or1'_parts_spies [dest]: "or1' A' A B NA X \ set evs +\ X \ parts (spies evs)" by blast -lemma or2_parts_spies [dest]: "or2 A B NA NB X:set evs -==> X:parts (spies evs)" +lemma or2_parts_spies [dest]: "or2 A B NA NB X \ set evs +\ X \ parts (spies evs)" by blast -lemma or3_parts_spies [dest]: "Says S B \NA, Y, Ciph B \NB, K\\:set evs -==> K:parts (spies evs)" +lemma or3_parts_spies [dest]: "Says S B \NA, Y, Ciph B \NB, K\\ \ set evs +\ K \ parts (spies evs)" by blast lemma or_is_regular [iff]: "regular or" @@ -119,8 +119,8 @@ subsection\guardedness of KAB\ -lemma Guard_KAB [rule_format]: "[| evs:or; A ~:bad; B ~:bad |] ==> -or3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" +lemma Guard_KAB [rule_format]: "[| evs \ or; A \ bad; B \ bad |] ==> +or3 A B NA NB K \ set evs \ GuardK K {shrK A,shrK B} (spies evs)" apply (erule or.induct) (* Nil *) apply simp_all @@ -140,8 +140,8 @@ subsection\guardedness of NB\ -lemma Guard_NB [rule_format]: "[| evs:or; B ~:bad |] ==> -or2 A B NA NB X:set evs --> Guard NB {shrK B} (spies evs)" +lemma Guard_NB [rule_format]: "[| evs \ or; B \ bad |] ==> +or2 A B NA NB X \ set evs \ Guard NB {shrK B} (spies evs)" apply (erule or.induct) (* Nil *) apply simp_all diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard_Public.thy --- a/src/HOL/Auth/Guard/Guard_Public.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard_Public.thy Thu Feb 15 13:04:36 2018 +0100 @@ -22,7 +22,7 @@ subsubsection\agent associated to a key\ definition agt :: "key => agent" where -"agt K == @A. K = priK A | K = pubK A" +"agt K == SOME A. K = priK A | K = pubK A" lemma agt_priK [simp]: "agt (priK A) = A" by (simp add: agt_def) @@ -32,18 +32,18 @@ subsubsection\basic facts about @{term initState}\ -lemma no_Crypt_in_parts_init [simp]: "Crypt K X ~:parts (initState A)" +lemma no_Crypt_in_parts_init [simp]: "Crypt K X \ parts (initState A)" by (cases A, auto simp: initState.simps) -lemma no_Crypt_in_analz_init [simp]: "Crypt K X ~:analz (initState A)" +lemma no_Crypt_in_analz_init [simp]: "Crypt K X \ analz (initState A)" by auto -lemma no_priK_in_analz_init [simp]: "A ~:bad -==> Key (priK A) ~:analz (initState Spy)" +lemma no_priK_in_analz_init [simp]: "A \ bad +\ Key (priK A) \ analz (initState Spy)" by (auto simp: initState.simps) -lemma priK_notin_initState_Friend [simp]: "A ~= Friend C -==> Key (priK A) ~: parts (initState (Friend C))" +lemma priK_notin_initState_Friend [simp]: "A \ Friend C +\ Key (priK A) \ parts (initState (Friend C))" by (auto simp: initState.simps) lemma keyset_init [iff]: "keyset (initState A)" @@ -52,9 +52,9 @@ subsubsection\sets of private keys\ definition priK_set :: "key set => bool" where -"priK_set Ks == ALL K. K:Ks --> (EX A. K = priK A)" +"priK_set Ks \ \K. K \ Ks \ (\A. K = priK A)" -lemma in_priK_set: "[| priK_set Ks; K:Ks |] ==> EX A. K = priK A" +lemma in_priK_set: "[| priK_set Ks; K \ Ks |] ==> \A. K = priK A" by (simp add: priK_set_def) lemma priK_set1 [iff]: "priK_set {priK A}" @@ -66,15 +66,15 @@ subsubsection\sets of good keys\ definition good :: "key set => bool" where -"good Ks == ALL K. K:Ks --> agt K ~:bad" +"good Ks == \K. K \ Ks \ agt K \ bad" -lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad" +lemma in_good: "[| good Ks; K \ Ks |] ==> agt K \ bad" by (simp add: good_def) -lemma good1 [simp]: "A ~:bad ==> good {priK A}" +lemma good1 [simp]: "A \ bad \ good {priK A}" by (simp add: good_def) -lemma good2 [simp]: "[| A ~:bad; B ~:bad |] ==> good {priK A, priK B}" +lemma good2 [simp]: "[| A \ bad; B \ bad |] ==> good {priK A, priK B}" by (simp add: good_def) subsubsection\greatest nonce used in a trace, 0 if there is no nonce\ @@ -84,7 +84,7 @@ "greatest [] = 0" | "greatest (ev # evs) = max (greatest_msg (msg ev)) (greatest evs)" -lemma greatest_is_greatest: "Nonce n:used evs ==> n <= greatest evs" +lemma greatest_is_greatest: "Nonce n \ used evs \ n \ greatest evs" apply (induct evs, auto simp: initState.simps) apply (drule used_sub_parts_used, safe) apply (drule greatest_msg_is_greatest, arith) @@ -92,10 +92,10 @@ subsubsection\function giving a new nonce\ -definition new :: "event list => nat" where -"new evs == Suc (greatest evs)" +definition new :: "event list \ nat" where +"new evs \ Suc (greatest evs)" -lemma new_isnt_used [iff]: "Nonce (new evs) ~:used evs" +lemma new_isnt_used [iff]: "Nonce (new evs) \ used evs" by (clarify, drule greatest_is_greatest, auto simp: new_def) subsection\Proofs About Guarded Messages\ @@ -109,13 +109,13 @@ lemmas invKey_invKey_substI = invKey [THEN ssubst] -lemma "Nonce n:parts {X} ==> Crypt (pubK A) X:guard n {priK A}" +lemma "Nonce n \ parts {X} \ Crypt (pubK A) X \ guard n {priK A}" apply (rule pubK_is_invKey_priK_substI, rule invKey_invKey_substI) by (rule Guard_Nonce, simp+) subsubsection\guardedness results\ -lemma sign_guard [intro]: "X:guard n Ks ==> sign A X:guard n Ks" +lemma sign_guard [intro]: "X \ guard n Ks \ sign A X \ guard n Ks" by (auto simp: sign_def) lemma Guard_init [iff]: "Guard n Ks (initState B)" @@ -125,38 +125,38 @@ ==> Guard n Ks (knows_max C evs)" by (simp add: knows_max_def) -lemma Nonce_not_used_Guard_spies [dest]: "Nonce n ~:used evs -==> Guard n Ks (spies evs)" +lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \ used evs +\ Guard n Ks (spies evs)" by (auto simp: Guard_def dest: not_used_not_known parts_sub) -lemma Nonce_not_used_Guard [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" by (auto simp: Guard_def dest: known_used parts_trans) -lemma Nonce_not_used_Guard_max [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard_max [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" by (auto simp: Guard_def dest: known_max_used parts_trans) -lemma Nonce_not_used_Guard_max' [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard_max' [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) by (auto simp: knows_max_def) subsubsection\regular protocols\ -definition regular :: "event list set => bool" where -"regular p == ALL evs A. evs:p --> (Key (priK A):parts (spies evs)) = (A:bad)" +definition regular :: "event list set \ bool" where +"regular p \ \evs A. evs \ p \ (Key (priK A) \ parts (spies evs)) = (A \ bad)" -lemma priK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==> -(Key (priK A):parts (spies evs)) = (A:bad)" +lemma priK_parts_iff_bad [simp]: "[| evs \ p; regular p |] ==> +(Key (priK A) \ parts (spies evs)) = (A \ bad)" by (auto simp: regular_def) -lemma priK_analz_iff_bad [simp]: "[| evs:p; regular p |] ==> -(Key (priK A):analz (spies evs)) = (A:bad)" +lemma priK_analz_iff_bad [simp]: "[| evs \ p; regular p |] ==> +(Key (priK A) \ analz (spies evs)) = (A \ bad)" by auto -lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs:p; -priK_set Ks; good Ks; regular p |] ==> Nonce n ~:analz (spies evs)" +lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \ p; +priK_set Ks; good Ks; regular p |] ==> Nonce n \ analz (spies evs)" apply (clarify, simp only: knows_decomp) apply (drule Guard_invKey_keyset, simp+, safe) apply (drule in_good, simp) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard_Shared.thy --- a/src/HOL/Auth/Guard/Guard_Shared.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard_Shared.thy Thu Feb 15 13:04:36 2018 +0100 @@ -20,25 +20,25 @@ subsubsection\agent associated to a key\ definition agt :: "key => agent" where -"agt K == @A. K = shrK A" +"agt K == SOME A. K = shrK A" lemma agt_shrK [simp]: "agt (shrK A) = A" by (simp add: agt_def) subsubsection\basic facts about @{term initState}\ -lemma no_Crypt_in_parts_init [simp]: "Crypt K X ~:parts (initState A)" +lemma no_Crypt_in_parts_init [simp]: "Crypt K X \ parts (initState A)" by (cases A, auto simp: initState.simps) -lemma no_Crypt_in_analz_init [simp]: "Crypt K X ~:analz (initState A)" +lemma no_Crypt_in_analz_init [simp]: "Crypt K X \ analz (initState A)" by auto -lemma no_shrK_in_analz_init [simp]: "A ~:bad -==> Key (shrK A) ~:analz (initState Spy)" +lemma no_shrK_in_analz_init [simp]: "A \ bad +\ Key (shrK A) \ analz (initState Spy)" by (auto simp: initState.simps) -lemma shrK_notin_initState_Friend [simp]: "A ~= Friend C -==> Key (shrK A) ~: parts (initState (Friend C))" +lemma shrK_notin_initState_Friend [simp]: "A \ Friend C +\ Key (shrK A) \ parts (initState (Friend C))" by (auto simp: initState.simps) lemma keyset_init [iff]: "keyset (initState A)" @@ -47,9 +47,9 @@ subsubsection\sets of symmetric keys\ definition shrK_set :: "key set => bool" where -"shrK_set Ks == ALL K. K:Ks --> (EX A. K = shrK A)" +"shrK_set Ks \ \K. K \ Ks \ (\A. K = shrK A)" -lemma in_shrK_set: "[| shrK_set Ks; K:Ks |] ==> EX A. K = shrK A" +lemma in_shrK_set: "[| shrK_set Ks; K \ Ks |] ==> \A. K = shrK A" by (simp add: shrK_set_def) lemma shrK_set1 [iff]: "shrK_set {shrK A}" @@ -60,16 +60,16 @@ subsubsection\sets of good keys\ -definition good :: "key set => bool" where -"good Ks == ALL K. K:Ks --> agt K ~:bad" +definition good :: "key set \ bool" where +"good Ks \ \K. K \ Ks \ agt K \ bad" -lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad" +lemma in_good: "[| good Ks; K \ Ks |] ==> agt K \ bad" by (simp add: good_def) -lemma good1 [simp]: "A ~:bad ==> good {shrK A}" +lemma good1 [simp]: "A \ bad \ good {shrK A}" by (simp add: good_def) -lemma good2 [simp]: "[| A ~:bad; B ~:bad |] ==> good {shrK A, shrK B}" +lemma good2 [simp]: "[| A \ bad; B \ bad |] ==> good {shrK A, shrK B}" by (simp add: good_def) @@ -84,16 +84,16 @@ lemmas invKey_invKey_substI = invKey [THEN ssubst] -lemma "Nonce n:parts {X} ==> Crypt (shrK A) X:guard n {shrK A}" +lemma "Nonce n \ parts {X} \ Crypt (shrK A) X \ guard n {shrK A}" apply (rule shrK_is_invKey_shrK_substI, rule invKey_invKey_substI) by (rule Guard_Nonce, simp+) subsubsection\guardedness results on nonces\ -lemma guard_ciph [simp]: "shrK A:Ks ==> Ciph A X:guard n Ks" +lemma guard_ciph [simp]: "shrK A \ Ks \ Ciph A X \ guard n Ks" by (rule Guard_Nonce, simp) -lemma guardK_ciph [simp]: "shrK A:Ks ==> Ciph A X:guardK n Ks" +lemma guardK_ciph [simp]: "shrK A \ Ks \ Ciph A X \ guardK n Ks" by (rule Guard_Key, simp) lemma Guard_init [iff]: "Guard n Ks (initState B)" @@ -103,45 +103,45 @@ ==> Guard n Ks (knows_max C evs)" by (simp add: knows_max_def) -lemma Nonce_not_used_Guard_spies [dest]: "Nonce n ~:used evs -==> Guard n Ks (spies evs)" +lemma Nonce_not_used_Guard_spies [dest]: "Nonce n \ used evs +\ Guard n Ks (spies evs)" by (auto simp: Guard_def dest: not_used_not_known parts_sub) -lemma Nonce_not_used_Guard [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows (Friend C) evs)" by (auto simp: Guard_def dest: known_used parts_trans) -lemma Nonce_not_used_Guard_max [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard_max [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows_max (Friend C) evs)" by (auto simp: Guard_def dest: known_max_used parts_trans) -lemma Nonce_not_used_Guard_max' [dest]: "[| evs:p; Nonce n ~:used evs; +lemma Nonce_not_used_Guard_max' [dest]: "[| evs \ p; Nonce n \ used evs; Gets_correct p; one_step p |] ==> Guard n Ks (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in Guard_mono) by (auto simp: knows_max_def) subsubsection\guardedness results on keys\ -lemma GuardK_init [simp]: "n ~:range shrK ==> GuardK n Ks (initState B)" +lemma GuardK_init [simp]: "n \ range shrK \ GuardK n Ks (initState B)" by (induct B, auto simp: GuardK_def initState.simps) -lemma GuardK_knows_max': "[| GuardK n A (knows_max' C evs); n ~:range shrK |] +lemma GuardK_knows_max': "[| GuardK n A (knows_max' C evs); n \ range shrK |] ==> GuardK n A (knows_max C evs)" by (simp add: knows_max_def) -lemma Key_not_used_GuardK_spies [dest]: "Key n ~:used evs -==> GuardK n A (spies evs)" +lemma Key_not_used_GuardK_spies [dest]: "Key n \ used evs +\ GuardK n A (spies evs)" by (auto simp: GuardK_def dest: not_used_not_known parts_sub) -lemma Key_not_used_GuardK [dest]: "[| evs:p; Key n ~:used evs; +lemma Key_not_used_GuardK [dest]: "[| evs \ p; Key n \ used evs; Gets_correct p; one_step p |] ==> GuardK n A (knows (Friend C) evs)" by (auto simp: GuardK_def dest: known_used parts_trans) -lemma Key_not_used_GuardK_max [dest]: "[| evs:p; Key n ~:used evs; +lemma Key_not_used_GuardK_max [dest]: "[| evs \ p; Key n \ used evs; Gets_correct p; one_step p |] ==> GuardK n A (knows_max (Friend C) evs)" by (auto simp: GuardK_def dest: known_max_used parts_trans) -lemma Key_not_used_GuardK_max' [dest]: "[| evs:p; Key n ~:used evs; +lemma Key_not_used_GuardK_max' [dest]: "[| evs \ p; Key n \ used evs; Gets_correct p; one_step p |] ==> GuardK n A (knows_max' (Friend C) evs)" apply (rule_tac H="knows_max (Friend C) evs" in GuardK_mono) by (auto simp: knows_max_def) @@ -149,18 +149,18 @@ subsubsection\regular protocols\ definition regular :: "event list set => bool" where -"regular p == ALL evs A. evs:p --> (Key (shrK A):parts (spies evs)) = (A:bad)" +"regular p \ \evs A. evs \ p \ (Key (shrK A) \ parts (spies evs)) = (A \ bad)" -lemma shrK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==> -(Key (shrK A):parts (spies evs)) = (A:bad)" +lemma shrK_parts_iff_bad [simp]: "[| evs \ p; regular p |] ==> +(Key (shrK A) \ parts (spies evs)) = (A \ bad)" by (auto simp: regular_def) -lemma shrK_analz_iff_bad [simp]: "[| evs:p; regular p |] ==> -(Key (shrK A):analz (spies evs)) = (A:bad)" +lemma shrK_analz_iff_bad [simp]: "[| evs \ p; regular p |] ==> +(Key (shrK A) \ analz (spies evs)) = (A \ bad)" by auto -lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs:p; -shrK_set Ks; good Ks; regular p |] ==> Nonce n ~:analz (spies evs)" +lemma Guard_Nonce_analz: "[| Guard n Ks (spies evs); evs \ p; +shrK_set Ks; good Ks; regular p |] ==> Nonce n \ analz (spies evs)" apply (clarify, simp only: knows_decomp) apply (drule Guard_invKey_keyset, simp+, safe) apply (drule in_good, simp) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Guard_Yahalom.thy --- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Thu Feb 15 13:04:36 2018 +0100 @@ -50,20 +50,20 @@ inductive_set ya :: "event list set" where - Nil: "[]:ya" + Nil: "[] \ ya" -| Fake: "[| evs:ya; X:synth (analz (spies evs)) |] ==> Says Spy B X # evs:ya" +| Fake: "[| evs \ ya; X \ synth (analz (spies evs)) |] ==> Says Spy B X # evs \ ya" -| YA1: "[| evs1:ya; Nonce NA ~:used evs1 |] ==> ya1 A B NA # evs1:ya" +| YA1: "[| evs1 \ ya; Nonce NA \ used evs1 |] ==> ya1 A B NA # evs1 \ ya" -| YA2: "[| evs2:ya; ya1' A' A B NA:set evs2; Nonce NB ~:used evs2 |] - ==> ya2 A B NA NB # evs2:ya" +| YA2: "[| evs2 \ ya; ya1' A' A B NA \ set evs2; Nonce NB \ used evs2 |] + ==> ya2 A B NA NB # evs2 \ ya" -| YA3: "[| evs3:ya; ya2' B' A B NA NB:set evs3; Key K ~:used evs3 |] - ==> ya3 A B NA NB K # evs3:ya" +| YA3: "[| evs3 \ ya; ya2' B' A B NA NB \ set evs3; Key K \ used evs3 |] + ==> ya3 A B NA NB K # evs3 \ ya" -| YA4: "[| evs4:ya; ya1 A B NA:set evs4; ya3' S Y A B NA NB K:set evs4 |] - ==> ya4 A B K NB Y # evs4:ya" +| YA4: "[| evs4 \ ya; ya1 A B NA \ set evs4; ya3' S Y A B NA NB K \ set evs4 |] + ==> ya4 A B K NB Y # evs4 \ ya" subsection\declarations for tactics\ @@ -73,17 +73,17 @@ subsection\general properties of ya\ -lemma ya_has_no_Gets: "evs:ya ==> ALL A X. Gets A X ~:set evs" +lemma ya_has_no_Gets: "evs \ ya \ \A X. Gets A X \ set evs" by (erule ya.induct, auto) lemma ya_is_Gets_correct [iff]: "Gets_correct ya" by (auto simp: Gets_correct_def dest: ya_has_no_Gets) lemma ya_is_one_step [iff]: "one_step ya" -by (unfold one_step_def, clarify, ind_cases "ev#evs:ya" for ev evs, auto) +by (unfold one_step_def, clarify, ind_cases "ev#evs \ ya" for ev evs, auto) -lemma ya_has_only_Says' [rule_format]: "evs:ya ==> -ev:set evs --> (EX A B X. ev=Says A B X)" +lemma ya_has_only_Says' [rule_format]: "evs \ ya \ +ev \ set evs \ (\A B X. ev=Says A B X)" by (erule ya.induct, auto) lemma ya_has_only_Says [iff]: "has_only_Says ya" @@ -96,8 +96,8 @@ subsection\guardedness of KAB\ -lemma Guard_KAB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==> -ya3 A B NA NB K:set evs --> GuardK K {shrK A,shrK B} (spies evs)" +lemma Guard_KAB [rule_format]: "[| evs \ ya; A \ bad; B \ bad |] ==> +ya3 A B NA NB K \ set evs \ GuardK K {shrK A,shrK B} (spies evs)" apply (erule ya.induct) (* Nil *) apply simp_all @@ -117,55 +117,55 @@ subsection\session keys are not symmetric keys\ -lemma KAB_isnt_shrK [rule_format]: "evs:ya ==> -ya3 A B NA NB K:set evs --> K ~:range shrK" +lemma KAB_isnt_shrK [rule_format]: "evs \ ya \ +ya3 A B NA NB K \ set evs \ K \ range shrK" by (erule ya.induct, auto) -lemma ya3_shrK: "evs:ya ==> ya3 A B NA NB (shrK C) ~:set evs" +lemma ya3_shrK: "evs \ ya \ ya3 A B NA NB (shrK C) \ set evs" by (blast dest: KAB_isnt_shrK) subsection\ya2' implies ya1'\ lemma ya2'_parts_imp_ya1'_parts [rule_format]: - "[| evs:ya; B ~:bad |] ==> - Ciph B \Agent A, Nonce NA, Nonce NB\:parts (spies evs) --> - \Agent A, Nonce NA\:spies evs" + "[| evs \ ya; B \ bad |] ==> + Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs) \ + \Agent A, Nonce NA\ \ spies evs" by (erule ya.induct, auto dest: Says_imp_spies intro: parts_parts) -lemma ya2'_imp_ya1'_parts: "[| ya2' B' A B NA NB:set evs; evs:ya; B ~:bad |] -==> \Agent A, Nonce NA\:spies evs" +lemma ya2'_imp_ya1'_parts: "[| ya2' B' A B NA NB \ set evs; evs \ ya; B \ bad |] +==> \Agent A, Nonce NA\ \ spies evs" by (blast dest: Says_imp_spies ya2'_parts_imp_ya1'_parts) subsection\uniqueness of NB\ -lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs:ya; B ~:bad; B' ~:bad |] ==> -Ciph B \Agent A, Nonce NA, Nonce NB\:parts (spies evs) --> -Ciph B' \Agent A', Nonce NA', Nonce NB\:parts (spies evs) --> -A=A' & B=B' & NA=NA'" +lemma NB_is_uniq_in_ya2'_parts [rule_format]: "[| evs \ ya; B \ bad; B' \ bad |] ==> +Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs) \ +Ciph B' \Agent A', Nonce NA', Nonce NB\ \ parts (spies evs) \ +A=A' \ B=B' \ NA=NA'" apply (erule ya.induct, simp_all, clarify) apply (drule Crypt_synth_insert, simp+) apply (drule Crypt_synth_insert, simp+, safe) apply (drule not_used_parts_false, simp+)+ by (drule Says_not_parts, simp+)+ -lemma NB_is_uniq_in_ya2': "[| ya2' C A B NA NB:set evs; -ya2' C' A' B' NA' NB:set evs; evs:ya; B ~:bad; B' ~:bad |] -==> A=A' & B=B' & NA=NA'" +lemma NB_is_uniq_in_ya2': "[| ya2' C A B NA NB \ set evs; +ya2' C' A' B' NA' NB \ set evs; evs \ ya; B \ bad; B' \ bad |] +==> A=A' \ B=B' \ NA=NA'" by (drule NB_is_uniq_in_ya2'_parts, auto dest: Says_imp_spies) subsection\ya3' implies ya2'\ -lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs:ya; A ~:bad |] ==> -Ciph A \Agent B, Key K, Nonce NA, Nonce NB\:parts (spies evs) ---> Ciph B \Agent A, Nonce NA, Nonce NB\:parts (spies evs)" +lemma ya3'_parts_imp_ya2'_parts [rule_format]: "[| evs \ ya; A \ bad |] ==> +Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts (spies evs) +\ Ciph B \Agent A, Nonce NA, Nonce NB\ \ parts (spies evs)" apply (erule ya.induct, simp_all) apply (clarify, drule Crypt_synth_insert, simp+) apply (blast intro: parts_sub, blast) by (auto dest: Says_imp_spies parts_parts) -lemma ya3'_parts_imp_ya2' [rule_format]: "[| evs:ya; A ~:bad |] ==> -Ciph A \Agent B, Key K, Nonce NA, Nonce NB\:parts (spies evs) ---> (EX B'. ya2' B' A B NA NB:set evs)" +lemma ya3'_parts_imp_ya2' [rule_format]: "[| evs \ ya; A \ bad |] ==> +Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts (spies evs) +\ (\B'. ya2' B' A B NA NB \ set evs)" apply (erule ya.induct, simp_all, safe) apply (drule Crypt_synth_insert, simp+) apply (drule Crypt_synth_insert, simp+, blast) @@ -173,30 +173,30 @@ apply blast by (auto dest: Says_imp_spies2 parts_parts) -lemma ya3'_imp_ya2': "[| ya3' S Y A B NA NB K:set evs; evs:ya; A ~:bad |] -==> (EX B'. ya2' B' A B NA NB:set evs)" +lemma ya3'_imp_ya2': "[| ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad |] +==> (\B'. ya2' B' A B NA NB \ set evs)" by (drule ya3'_parts_imp_ya2', auto dest: Says_imp_spies) subsection\ya3' implies ya3\ -lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs:ya; A ~:bad |] ==> -Ciph A \Agent B, Key K, Nonce NA, Nonce NB\:parts(spies evs) ---> ya3 A B NA NB K:set evs" +lemma ya3'_parts_imp_ya3 [rule_format]: "[| evs \ ya; A \ bad |] ==> +Ciph A \Agent B, Key K, Nonce NA, Nonce NB\ \ parts(spies evs) +\ ya3 A B NA NB K \ set evs" apply (erule ya.induct, simp_all, safe) apply (drule Crypt_synth_insert, simp+) by (blast dest: Says_imp_spies2 parts_parts) -lemma ya3'_imp_ya3: "[| ya3' S Y A B NA NB K:set evs; evs:ya; A ~:bad |] -==> ya3 A B NA NB K:set evs" +lemma ya3'_imp_ya3: "[| ya3' S Y A B NA NB K \ set evs; evs \ ya; A \ bad |] +==> ya3 A B NA NB K \ set evs" by (blast dest: Says_imp_spies ya3'_parts_imp_ya3) subsection\guardedness of NB\ -definition ya_keys :: "agent => agent => nat => nat => event list => key set" where -"ya_keys A B NA NB evs == {shrK A,shrK B} Un {K. ya3 A B NA NB K:set evs}" +definition ya_keys :: "agent \ agent \ nat \ nat \ event list \ key set" where +"ya_keys A B NA NB evs \ {shrK A,shrK B} \ {K. ya3 A B NA NB K \ set evs}" -lemma Guard_NB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==> -ya2 A B NA NB:set evs --> Guard NB (ya_keys A B NA NB evs) (spies evs)" +lemma Guard_NB [rule_format]: "[| evs \ ya; A \ bad; B \ bad |] ==> +ya2 A B NA NB \ set evs \ Guard NB (ya_keys A B NA NB evs) (spies evs)" apply (erule ya.induct) (* Nil *) apply (simp_all add: ya_keys_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/List_Msg.thy --- a/src/HOL/Auth/Guard/List_Msg.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/List_Msg.thy Thu Feb 15 13:04:36 2018 +0100 @@ -37,7 +37,7 @@ "len (cons x l) = Suc (len l)" | "len other = 0" -lemma len_not_empty: "n < len l ==> EX x l'. l = cons x l'" +lemma len_not_empty: "n < len l \ \x l'. l = cons x l'" by (cases l) auto subsubsection\membership\ @@ -113,36 +113,36 @@ inductive_set agl :: "msg set" where - Nil[intro]: "nil:agl" -| Cons[intro]: "[| A:agent; I:agl |] ==> cons (Agent A) I :agl" + Nil[intro]: "nil \ agl" +| Cons[intro]: "[| A \ agent; I \ agl |] ==> cons (Agent A) I \ agl" subsubsection\basic facts about agent lists\ -lemma del_in_agl [intro]: "I:agl ==> del (a,I):agl" +lemma del_in_agl [intro]: "I \ agl \ del (a,I) \ agl" by (erule agl.induct, auto) -lemma app_in_agl [intro]: "[| I:agl; J:agl |] ==> app (I,J):agl" +lemma app_in_agl [intro]: "[| I \ agl; J \ agl |] ==> app (I,J) \ agl" by (erule agl.induct, auto) -lemma no_Key_in_agl: "I:agl ==> Key K ~:parts {I}" +lemma no_Key_in_agl: "I \ agl \ Key K \ parts {I}" by (erule agl.induct, auto) -lemma no_Nonce_in_agl: "I:agl ==> Nonce n ~:parts {I}" +lemma no_Nonce_in_agl: "I \ agl \ Nonce n \ parts {I}" by (erule agl.induct, auto) -lemma no_Key_in_appdel: "[| I:agl; J:agl |] ==> -Key K ~:parts {app (J, del (Agent B, I))}" +lemma no_Key_in_appdel: "[| I \ agl; J \ agl |] ==> +Key K \ parts {app (J, del (Agent B, I))}" by (rule no_Key_in_agl, auto) -lemma no_Nonce_in_appdel: "[| I:agl; J:agl |] ==> -Nonce n ~:parts {app (J, del (Agent B, I))}" +lemma no_Nonce_in_appdel: "[| I \ agl; J \ agl |] ==> +Nonce n \ parts {app (J, del (Agent B, I))}" by (rule no_Nonce_in_agl, auto) -lemma no_Crypt_in_agl: "I:agl ==> Crypt K X ~:parts {I}" +lemma no_Crypt_in_agl: "I \ agl \ Crypt K X \ parts {I}" by (erule agl.induct, auto) -lemma no_Crypt_in_appdel: "[| I:agl; J:agl |] ==> -Crypt K X ~:parts {app (J, del (Agent B,I))}" +lemma no_Crypt_in_appdel: "[| I \ agl; J \ agl |] ==> +Crypt K X \ parts {app (J, del (Agent B,I))}" by (rule no_Crypt_in_agl, auto) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/P1.thy --- a/src/HOL/Auth/Guard/P1.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/P1.thy Thu Feb 15 13:04:36 2018 +0100 @@ -45,7 +45,7 @@ = (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')" by (auto simp: chain_def Let_def) -lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}" +lemma Nonce_in_chain [iff]: "Nonce ofr \ parts {chain B ofr A L C}" by (auto simp: chain_def sign_def) subsubsection\agent whose key is used to sign an offer\ @@ -81,7 +81,7 @@ = (A=A' & n=n' & B=B')" by (auto simp: anchor_def) -lemma Nonce_in_anchor [iff]: "Nonce n:parts {anchor A n B}" +lemma Nonce_in_anchor [iff]: "Nonce n \ parts {anchor A n B}" by (auto simp: anchor_def) lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A" @@ -103,7 +103,7 @@ = (A=A' & r=r' & n=n' & I=I' & B=B')" by (auto simp: reqm_def) -lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}" +lemma Nonce_in_reqm [iff]: "Nonce n \ parts {reqm A r n I B}" by (auto simp: reqm_def) definition req :: "agent => nat => nat => msg => agent => event" where @@ -125,7 +125,7 @@ ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: prom_def) -lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}" +lemma Nonce_in_prom [iff]: "Nonce ofr \ parts {prom B ofr A r I L J C}" by (auto simp: prom_def) definition pro :: "agent => nat => agent => nat => msg => msg => @@ -141,21 +141,21 @@ inductive_set p1 :: "event list set" where - Nil: "[]:p1" + Nil: "[] \ p1" -| Fake: "[| evsf:p1; X:synth (analz (spies evsf)) |] ==> Says Spy B X # evsf : p1" +| Fake: "[| evsf \ p1; X \ synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \ p1" -| Request: "[| evsr:p1; Nonce n ~:used evsr; I:agl |] ==> req A r n I B # evsr : p1" +| Request: "[| evsr \ p1; Nonce n \ used evsr; I \ agl |] ==> req A r n I B # evsr \ p1" -| Propose: "[| evsp:p1; Says A' B \Agent A,Number r,I,cons M L\:set evsp; - I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I))); - Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p1" +| Propose: "[| evsp \ p1; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; + I \ agl; J \ agl; isin (Agent C, app (J, del (Agent B, I))); + Nonce ofr \ used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \ p1" subsubsection\Composition of Traces\ -lemma "evs':p1 ==> - evs:p1 & (ALL n. Nonce n:used evs' --> Nonce n ~:used evs) --> - evs'@evs : p1" +lemma "evs' \ p1 \ + evs \ p1 \ (\n. Nonce n \ used evs' \ Nonce n \ used evs) \ + evs' @ evs \ p1" apply (erule p1.induct, safe) apply (simp_all add: used_ConsI) apply (erule p1.Fake, erule synth_sub, rule analz_mono, rule knows_sub_app) @@ -168,30 +168,30 @@ subsubsection\Valid Offer Lists\ inductive_set - valid :: "agent => nat => agent => msg set" + valid :: "agent \ nat \ agent \ msg set" for A :: agent and n :: nat and B :: agent where - Request [intro]: "cons (anchor A n B) nil:valid A n B" + Request [intro]: "cons (anchor A n B) nil \ valid A n B" -| Propose [intro]: "L:valid A n B -==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B" +| Propose [intro]: "L \ valid A n B +\ cons (chain (next_shop (head L)) ofr A L C) L \ valid A n B" subsubsection\basic properties of valid\ -lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'" +lemma valid_not_empty: "L \ valid A n B \ \M L'. L = cons M L'" by (erule valid.cases, auto) -lemma valid_pos_len: "L:valid A n B ==> 0 < len L" +lemma valid_pos_len: "L \ valid A n B \ 0 < len L" by (erule valid.induct, auto) subsubsection\offers of an offer list\ -definition offer_nonces :: "msg => msg set" where -"offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}" +definition offer_nonces :: "msg \ msg set" where +"offer_nonces L \ {X. X \ parts {L} \ (\n. X = Nonce n)}" subsubsection\the originator can get the offers\ -lemma "L:valid A n B ==> offer_nonces L <= analz (insert L (initState A))" +lemma "L \ valid A n B \ offer_nonces L \ analz (insert L (initState A))" by (erule valid.induct, auto simp: anchor_def chain_def sign_def offer_nonces_def initState.simps) @@ -207,22 +207,22 @@ "shops (cons M L) = cons (shop M) (shops L)" | "shops other = other" -lemma shops_in_agl: "L:valid A n B ==> shops L:agl" +lemma shops_in_agl: "L \ valid A n B \ shops L \ agl" by (erule valid.induct, auto simp: anchor_def chain_def sign_def) subsubsection\builds a trace from an itinerary\ -fun offer_list :: "agent * nat * agent * msg * nat => msg" where +fun offer_list :: "agent \ nat \ agent \ msg \ nat \ msg" where "offer_list (A,n,B,nil,ofr) = cons (anchor A n B) nil" | "offer_list (A,n,B,cons (Agent C) I,ofr) = ( let L = offer_list (A,n,B,I,Suc ofr) in cons (chain (next_shop (head L)) ofr A L C) L)" -lemma "I:agl ==> ALL ofr. offer_list (A,n,B,I,ofr):valid A n B" +lemma "I \ agl \ \ofr. offer_list (A,n,B,I,ofr) \ valid A n B" by (erule agl.induct, auto) -fun trace :: "agent * nat * agent * nat * msg * msg * msg -=> event list" where +fun trace :: "agent \ nat \ agent \ nat \ msg \ msg \ msg +\ event list" where "trace (B,ofr,A,r,I,L,nil) = []" | "trace (B,ofr,A,r,I,L,cons (Agent D) K) = ( let C = (if K=nil then B else agt_nb (head K)) in @@ -232,8 +232,8 @@ pro C (Suc ofr) A r I' L nil D # trace (B,Suc ofr,A,r,I'',tail L,K))" -definition trace' :: "agent => nat => nat => msg => agent => nat => event list" where -"trace' A r n I B ofr == ( +definition trace' :: "agent \ nat \ nat \ msg \ agent \ nat \ event list" where +"trace' A r n I B ofr \ ( let AI = cons (Agent A) I in let L = offer_list (A,n,B,AI,ofr) in trace (B,ofr,A,r,nil,L,AI))" @@ -242,8 +242,8 @@ subsubsection\there is a trace in which the originator receives a valid answer\ -lemma p1_not_empty: "evs:p1 ==> req A r n I B:set evs --> -(EX evs'. evs'@evs:p1 & pro B' ofr A r I' L J A:set evs' & L:valid A n B)" +lemma p1_not_empty: "evs \ p1 \ req A r n I B \ set evs \ +(\evs'. evs' @ evs \ p1 \ pro B' ofr A r I' L J A \ set evs' \ L \ valid A n B)" oops @@ -255,66 +255,66 @@ subsubsection\strong forward integrity: except the last one, no offer can be modified\ -lemma strong_forward_integrity: "ALL L. Suc i < len L ---> L:valid A n B & repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)" +lemma strong_forward_integrity: "\L. Suc i < len L +\ L \ valid A n B \ repl (L,Suc i,M) \ valid A n B \ M = ith (L,Suc i)" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,xa,l'a\:valid A n B" for x xa l'a) -apply (ind_cases "\x,M,l'a\:valid A n B" for x l'a) +apply (ind_cases "\x,xa,l'a\ \ valid A n B" for x xa l'a) +apply (ind_cases "\x,M,l'a\ \ valid A n B" for x l'a) apply (simp add: chain_def) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,repl(l',Suc na,M)\:valid A n B" for x l' na) +apply (ind_cases "\x,repl(l',Suc na,M)\ \ valid A n B" for x l' na) apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') by (drule_tac x=l' in spec, simp, blast) subsubsection\insertion resilience: except at the beginning, no offer can be inserted\ -lemma chain_isnt_head [simp]: "L:valid A n B ==> -head L ~= chain (next_shop (head L)) ofr A L C" +lemma chain_isnt_head [simp]: "L \ valid A n B \ +head L \ chain (next_shop (head L)) ofr A L C" by (erule valid.induct, auto simp: chain_def sign_def anchor_def) -lemma insertion_resilience: "ALL L. L:valid A n B --> Suc i < len L ---> ins (L,Suc i,M) ~:valid A n B" +lemma insertion_resilience: "\L. L \ valid A n B \ Suc i < len L +\ ins (L,Suc i,M) \ valid A n B" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l', simp) -apply (ind_cases "\x,M,l'\:valid A n B" for x l', clarsimp) -apply (ind_cases "\head l',l'\:valid A n B" for l', simp, simp) +apply (ind_cases "\x,l'\ \ valid A n B" for x l', simp) +apply (ind_cases "\x,M,l'\ \ valid A n B" for x l', clarsimp) +apply (ind_cases "\head l',l'\ \ valid A n B" for l', simp, simp) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,ins(l',Suc na,M)\:valid A n B" for x l' na) +apply (ind_cases "\x,ins(l',Suc na,M)\ \ valid A n B" for x l' na) apply (frule len_not_empty, clarsimp) by (drule_tac x=l' in spec, clarsimp) subsubsection\truncation resilience: only shop i can truncate at offer i\ -lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L ---> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))" +lemma truncation_resilience: "\L. L \ valid A n B \ Suc i < len L +\ cons M (trunc (L,Suc i)) \ valid A n B \ shop M = shop (ith (L,i))" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) -apply (ind_cases "\M,l'\:valid A n B" for l') +apply (ind_cases "\M,l'\ \ valid A n B" for l') apply (frule len_not_empty, clarsimp, simp) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) by (drule_tac x=l' in spec, clarsimp) @@ -326,37 +326,37 @@ subsubsection\get components of a message\ -lemma get_ML [dest]: "Says A' B \A,r,I,M,L\:set evs ==> -M:parts (spies evs) & L:parts (spies evs)" +lemma get_ML [dest]: "Says A' B \A,r,I,M,L\ \ set evs \ +M \ parts (spies evs) \ L \ parts (spies evs)" by blast subsubsection\general properties of p1\ lemma reqm_neq_prom [iff]: -"reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C" +"reqm A r n I B \ prom B' ofr A' r' I' (cons M L) J C" by (auto simp: reqm_def prom_def) lemma prom_neq_reqm [iff]: -"prom B' ofr A' r' I' (cons M L) J C ~= reqm A r n I B" +"prom B' ofr A' r' I' (cons M L) J C \ reqm A r n I B" by (auto simp: reqm_def prom_def) -lemma req_neq_pro [iff]: "req A r n I B ~= pro B' ofr A' r' I' (cons M L) J C" +lemma req_neq_pro [iff]: "req A r n I B \ pro B' ofr A' r' I' (cons M L) J C" by (auto simp: req_def pro_def) -lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C ~= req A r n I B" +lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C \ req A r n I B" by (auto simp: req_def pro_def) -lemma p1_has_no_Gets: "evs:p1 ==> ALL A X. Gets A X ~:set evs" +lemma p1_has_no_Gets: "evs \ p1 \ \A X. Gets A X \ set evs" by (erule p1.induct, auto simp: req_def pro_def) lemma p1_is_Gets_correct [iff]: "Gets_correct p1" by (auto simp: Gets_correct_def dest: p1_has_no_Gets) lemma p1_is_one_step [iff]: "one_step p1" -by (unfold one_step_def, clarify, ind_cases "ev#evs:p1" for ev evs, auto) +by (unfold one_step_def, clarify, ind_cases "ev#evs \ p1" for ev evs, auto) -lemma p1_has_only_Says' [rule_format]: "evs:p1 ==> -ev:set evs --> (EX A B X. ev=Says A B X)" +lemma p1_has_only_Says' [rule_format]: "evs \ p1 \ +ev \ set evs \ (\A B X. ev=Says A B X)" by (erule p1.induct, auto simp: req_def pro_def) lemma p1_has_only_Says [iff]: "has_only_Says p1" @@ -372,8 +372,8 @@ subsubsection\private keys are safe\ lemma priK_parts_Friend_imp_bad [rule_format,dest]: - "[| evs:p1; Friend B ~= A |] - ==> (Key (priK A):parts (knows (Friend B) evs)) --> (A:bad)" + "[| evs \ p1; Friend B \ A |] + ==> (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" apply (erule p1.induct) apply (simp_all add: initState.simps knows.simps pro_def prom_def req_def reqm_def anchor_def chain_def sign_def) @@ -383,12 +383,12 @@ done lemma priK_analz_Friend_imp_bad [rule_format,dest]: - "[| evs:p1; Friend B ~= A |] -==> (Key (priK A):analz (knows (Friend B) evs)) --> (A:bad)" + "[| evs \ p1; Friend B \ A |] +==> (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" by auto -lemma priK_notin_knows_max_Friend: "[| evs:p1; A ~:bad; A ~= Friend C |] -==> Key (priK A) ~:analz (knows_max (Friend C) evs)" +lemma priK_notin_knows_max_Friend: "[| evs \ p1; A \ bad; A \ Friend C |] +==> Key (priK A) \ analz (knows_max (Friend C) evs)" apply (rule not_parts_not_analz, simp add: knows_max_def, safe) apply (drule_tac H="spies' evs" in parts_sub) apply (rule_tac p=p1 in knows_max'_sub_spies', simp+) @@ -397,78 +397,78 @@ subsubsection\general guardedness properties\ -lemma agl_guard [intro]: "I:agl ==> I:guard n Ks" +lemma agl_guard [intro]: "I \ agl \ I \ guard n Ks" by (erule agl.induct, auto) -lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\:set evs; -Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" +lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" by (auto dest: Says_to_knows_max') -lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\:set evs; -Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" +lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" by (auto dest: Says_from_knows_max') -lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\:set evs; -Nonce n ~:used evs |] ==> L:guard n Ks" +lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\ \ set evs; +Nonce n \ used evs |] ==> L \ guard n Ks" by (drule not_used_not_parts, auto) subsubsection\guardedness of messages\ -lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}" +lemma chain_guard [iff]: "chain B ofr A L C \ guard n {priK A}" by (case_tac "ofr=n", auto simp: chain_def sign_def) -lemma chain_guard_Nonce_neq [intro]: "n ~= ofr -==> chain B ofr A' L C:guard n {priK A}" +lemma chain_guard_Nonce_neq [intro]: "n \ ofr +\ chain B ofr A' L C \ guard n {priK A}" by (auto simp: chain_def sign_def) -lemma anchor_guard [iff]: "anchor A n' B:guard n {priK A}" +lemma anchor_guard [iff]: "anchor A n' B \ guard n {priK A}" by (case_tac "n'=n", auto simp: anchor_def) -lemma anchor_guard_Nonce_neq [intro]: "n ~= n' -==> anchor A' n' B:guard n {priK A}" +lemma anchor_guard_Nonce_neq [intro]: "n \ n' +\ anchor A' n' B \ guard n {priK A}" by (auto simp: anchor_def) -lemma reqm_guard [intro]: "I:agl ==> reqm A r n' I B:guard n {priK A}" +lemma reqm_guard [intro]: "I \ agl \ reqm A r n' I B \ guard n {priK A}" by (case_tac "n'=n", auto simp: reqm_def) -lemma reqm_guard_Nonce_neq [intro]: "[| n ~= n'; I:agl |] -==> reqm A' r n' I B:guard n {priK A}" +lemma reqm_guard_Nonce_neq [intro]: "[| n \ n'; I \ agl |] +==> reqm A' r n' I B \ guard n {priK A}" by (auto simp: reqm_def) -lemma prom_guard [intro]: "[| I:agl; J:agl; L:guard n {priK A} |] -==> prom B ofr A r I L J C:guard n {priK A}" +lemma prom_guard [intro]: "[| I \ agl; J \ agl; L \ guard n {priK A} |] +==> prom B ofr A r I L J C \ guard n {priK A}" by (auto simp: prom_def) -lemma prom_guard_Nonce_neq [intro]: "[| n ~= ofr; I:agl; J:agl; -L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}" +lemma prom_guard_Nonce_neq [intro]: "[| n \ ofr; I \ agl; J \ agl; +L \ guard n {priK A} |] ==> prom B ofr A' r I L J C \ guard n {priK A}" by (auto simp: prom_def) subsubsection\Nonce uniqueness\ -lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr" +lemma uniq_Nonce_in_chain [dest]: "Nonce k \ parts {chain B ofr A L C} \ k=ofr" by (auto simp: chain_def sign_def) -lemma uniq_Nonce_in_anchor [dest]: "Nonce k:parts {anchor A n B} ==> k=n" +lemma uniq_Nonce_in_anchor [dest]: "Nonce k \ parts {anchor A n B} \ k=n" by (auto simp: anchor_def chain_def sign_def) -lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k:parts {reqm A r n I B}; -I:agl |] ==> k=n" +lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \ parts {reqm A r n I B}; +I \ agl |] ==> k=n" by (auto simp: reqm_def dest: no_Nonce_in_agl) -lemma uniq_Nonce_in_prom [dest]: "[| Nonce k:parts {prom B ofr A r I L J C}; -I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr" +lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \ parts {prom B ofr A r I L J C}; +I \ agl; J \ agl; Nonce k \ parts {L} |] ==> k=ofr" by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) subsubsection\requests are guarded\ -lemma req_imp_Guard [rule_format]: "[| evs:p1; A ~:bad |] ==> -req A r n I B:set evs --> Guard n {priK A} (spies evs)" +lemma req_imp_Guard [rule_format]: "[| evs \ p1; A \ bad |] ==> +req A r n I B \ set evs \ Guard n {priK A} (spies evs)" apply (erule p1.induct, simp) apply (simp add: req_def knows.simps, safe) apply (erule in_synth_Guard, erule Guard_analz, simp) by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) -lemma req_imp_Guard_Friend: "[| evs:p1; A ~:bad; req A r n I B:set evs |] +lemma req_imp_Guard_Friend: "[| evs \ p1; A \ bad; req A r n I B \ set evs |] ==> Guard n {priK A} (knows_max (Friend C) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) @@ -479,8 +479,8 @@ subsubsection\propositions are guarded\ -lemma pro_imp_Guard [rule_format]: "[| evs:p1; B ~:bad; A ~:bad |] ==> -pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)" +lemma pro_imp_Guard [rule_format]: "[| evs \ p1; B \ bad; A \ bad |] ==> +pro B ofr A r I (cons M L) J C \ set evs \ Guard ofr {priK A} (spies evs)" apply (erule p1.induct) (* +3 subgoals *) (* Nil *) apply simp @@ -516,8 +516,8 @@ apply (simp add: pro_def) by (blast dest: Says_imp_knows_Spy) -lemma pro_imp_Guard_Friend: "[| evs:p1; B ~:bad; A ~:bad; -pro B ofr A r I (cons M L) J C:set evs |] +lemma pro_imp_Guard_Friend: "[| evs \ p1; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs |] ==> Guard ofr {priK A} (knows_max (Friend D) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) @@ -529,23 +529,23 @@ subsubsection\data confidentiality: no one other than the originator can decrypt the offers\ -lemma Nonce_req_notin_spies: "[| evs:p1; req A r n I B:set evs; A ~:bad |] -==> Nonce n ~:analz (spies evs)" +lemma Nonce_req_notin_spies: "[| evs \ p1; req A r n I B \ set evs; A \ bad |] +==> Nonce n \ analz (spies evs)" by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_req_notin_knows_max_Friend: "[| evs:p1; req A r n I B:set evs; -A ~:bad; A ~= Friend C |] ==> Nonce n ~:analz (knows_max (Friend C) evs)" +lemma Nonce_req_notin_knows_max_Friend: "[| evs \ p1; req A r n I B \ set evs; +A \ bad; A \ Friend C |] ==> Nonce n \ analz (knows_max (Friend C) evs)" apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) -lemma Nonce_pro_notin_spies: "[| evs:p1; B ~:bad; A ~:bad; -pro B ofr A r I (cons M L) J C:set evs |] ==> Nonce ofr ~:analz (spies evs)" +lemma Nonce_pro_notin_spies: "[| evs \ p1; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs |] ==> Nonce ofr \ analz (spies evs)" by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_pro_notin_knows_max_Friend: "[| evs:p1; B ~:bad; A ~:bad; -A ~= Friend D; pro B ofr A r I (cons M L) J C:set evs |] -==> Nonce ofr ~:analz (knows_max (Friend D) evs)" +lemma Nonce_pro_notin_knows_max_Friend: "[| evs \ p1; B \ bad; A \ bad; +A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs |] +==> Nonce ofr \ analz (knows_max (Friend D) evs)" apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) @@ -553,59 +553,59 @@ subsubsection\non repudiability: an offer signed by B has been sent by B\ -lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'" +lemma Crypt_reqm: "[| Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl |] ==> A=A'" by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) -lemma Crypt_prom: "[| Crypt (priK A) X:parts {prom B ofr A' r I L J C}; -I:agl; J:agl |] ==> A=B | Crypt (priK A) X:parts {L}" +lemma Crypt_prom: "[| Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; +I \ agl; J \ agl |] ==> A=B \ Crypt (priK A) X \ parts {L}" apply (simp add: prom_def anchor_def chain_def sign_def) by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) -lemma Crypt_safeness: "[| evs:p1; A ~:bad |] ==> Crypt (priK A) X:parts (spies evs) ---> (EX B Y. Says A B Y:set evs & Crypt (priK A) X:parts {Y})" +lemma Crypt_safeness: "[| evs \ p1; A \ bad |] ==> Crypt (priK A) X \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs \ Crypt (priK A) X \ parts {Y})" apply (erule p1.induct) (* Nil *) apply simp (* Fake *) apply clarsimp -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (erule disjE) apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) (* Request *) apply (simp add: req_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (erule disjE) apply (frule Crypt_reqm, simp, clarify) apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast) (* Propose *) apply (simp add: pro_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (rotate_tac -1, erule disjE) apply (frule Crypt_prom, simp, simp) apply (rotate_tac -1, erule disjE) apply (rule_tac x=C in exI) apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast) -apply (subgoal_tac "cons M L:parts (spies evsp)") +apply (subgoal_tac "cons M L \ parts (spies evsp)") apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast) apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj) apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) by auto -lemma Crypt_Hash_imp_sign: "[| evs:p1; A ~:bad |] ==> -Crypt (priK A) (Hash X):parts (spies evs) ---> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" +lemma Crypt_Hash_imp_sign: "[| evs \ p1; A \ bad |] ==> +Crypt (priK A) (Hash X) \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (erule p1.induct) (* Nil *) apply simp (* Fake *) apply clarsimp -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (erule disjE) apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) (* Request *) apply (simp add: req_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (erule disjE) apply (frule Crypt_reqm, simp+) @@ -614,7 +614,7 @@ apply (simp add: chain_def sign_def, blast) (* Propose *) apply (simp add: pro_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (rotate_tac -1, erule disjE) apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel) @@ -628,8 +628,8 @@ apply (blast del: MPair_parts)+ done -lemma sign_safeness: "[| evs:p1; A ~:bad |] ==> sign A X:parts (spies evs) ---> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" +lemma sign_safeness: "[| evs \ p1; A \ bad |] ==> sign A X \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (clarify, simp add: sign_def, frule parts.Snd) apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/P2.thy --- a/src/HOL/Auth/Guard/P2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/P2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -32,7 +32,7 @@ = (B=B' & ofr=ofr' & A=A' & head L = head L' & C=C')" by (auto simp: chain_def Let_def) -lemma Nonce_in_chain [iff]: "Nonce ofr:parts {chain B ofr A L C}" +lemma Nonce_in_chain [iff]: "Nonce ofr \ parts {chain B ofr A L C}" by (auto simp: chain_def sign_def) subsubsection\agent whose key is used to sign an offer\ @@ -65,10 +65,10 @@ "anchor A n B == chain A n A (cons nil nil) B" lemma anchor_inj [iff]: - "(anchor A n B = anchor A' n' B') = (A=A' & n=n' & B=B')" + "(anchor A n B = anchor A' n' B') = (A=A' \ n=n' \ B=B')" by (auto simp: anchor_def) -lemma Nonce_in_anchor [iff]: "Nonce n:parts {anchor A n B}" +lemma Nonce_in_anchor [iff]: "Nonce n \ parts {anchor A n B}" by (auto simp: anchor_def) lemma shop_anchor [simp]: "shop (anchor A n B) = Agent A" @@ -84,7 +84,7 @@ = (A=A' & r=r' & n=n' & I=I' & B=B')" by (auto simp: reqm_def) -lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}" +lemma Nonce_in_reqm [iff]: "Nonce n \ parts {reqm A r n I B}" by (auto simp: reqm_def) definition req :: "agent => nat => nat => msg => agent => event" where @@ -105,7 +105,7 @@ ==> B=B' & ofr=ofr' & A=A' & r=r' & L=L' & C=C'" by (auto simp: prom_def) -lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}" +lemma Nonce_in_prom [iff]: "Nonce ofr \ parts {prom B ofr A r I L J C}" by (auto simp: prom_def) definition pro :: "agent => nat => agent => nat => msg => msg => @@ -121,38 +121,38 @@ inductive_set p2 :: "event list set" where - Nil: "[]:p2" + Nil: "[] \ p2" -| Fake: "[| evsf:p2; X:synth (analz (spies evsf)) |] ==> Says Spy B X # evsf : p2" +| Fake: "[| evsf \ p2; X \ synth (analz (spies evsf)) |] ==> Says Spy B X # evsf \ p2" -| Request: "[| evsr:p2; Nonce n ~:used evsr; I:agl |] ==> req A r n I B # evsr : p2" +| Request: "[| evsr \ p2; Nonce n \ used evsr; I \ agl |] ==> req A r n I B # evsr \ p2" -| Propose: "[| evsp:p2; Says A' B \Agent A,Number r,I,cons M L\:set evsp; - I:agl; J:agl; isin (Agent C, app (J, del (Agent B, I))); - Nonce ofr ~:used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp : p2" +| Propose: "[| evsp \ p2; Says A' B \Agent A,Number r,I,cons M L\ \ set evsp; + I \ agl; J \ agl; isin (Agent C, app (J, del (Agent B, I))); + Nonce ofr \ used evsp |] ==> pro B ofr A r I (cons M L) J C # evsp \ p2" subsubsection\valid offer lists\ inductive_set - valid :: "agent => nat => agent => msg set" + valid :: "agent \ nat \ agent \ msg set" for A :: agent and n :: nat and B :: agent where - Request [intro]: "cons (anchor A n B) nil:valid A n B" + Request [intro]: "cons (anchor A n B) nil \ valid A n B" -| Propose [intro]: "L:valid A n B - ==> cons (chain (next_shop (head L)) ofr A L C) L:valid A n B" +| Propose [intro]: "L \ valid A n B + \ cons (chain (next_shop (head L)) ofr A L C) L \ valid A n B" subsubsection\basic properties of valid\ -lemma valid_not_empty: "L:valid A n B ==> EX M L'. L = cons M L'" +lemma valid_not_empty: "L \ valid A n B \ \M L'. L = cons M L'" by (erule valid.cases, auto) -lemma valid_pos_len: "L:valid A n B ==> 0 < len L" +lemma valid_pos_len: "L \ valid A n B \ 0 < len L" by (erule valid.induct, auto) subsubsection\list of offers\ -fun offers :: "msg => msg" +fun offers :: "msg \ msg" where "offers (cons M L) = cons \shop M, nonce M\ (offers L)" | "offers other = nil" @@ -166,66 +166,66 @@ subsection\strong forward integrity: except the last one, no offer can be modified\ -lemma strong_forward_integrity: "ALL L. Suc i < len L ---> L:valid A n B --> repl (L,Suc i,M):valid A n B --> M = ith (L,Suc i)" +lemma strong_forward_integrity: "\L. Suc i < len L +\ L \ valid A n B \ repl (L,Suc i,M) \ valid A n B \ M = ith (L,Suc i)" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,xa,l'a\:valid A n B" for x xa l'a) -apply (ind_cases "\x,M,l'a\:valid A n B" for x l'a) +apply (ind_cases "\x,xa,l'a\ \ valid A n B" for x xa l'a) +apply (ind_cases "\x,M,l'a\ \ valid A n B" for x l'a) apply (simp add: chain_def) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,repl(l',Suc na,M)\:valid A n B" for x l' na) +apply (ind_cases "\x,repl(l',Suc na,M)\ \ valid A n B" for x l' na) apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') by (drule_tac x=l' in spec, simp, blast) subsection\insertion resilience: except at the beginning, no offer can be inserted\ -lemma chain_isnt_head [simp]: "L:valid A n B ==> -head L ~= chain (next_shop (head L)) ofr A L C" +lemma chain_isnt_head [simp]: "L \ valid A n B \ +head L \ chain (next_shop (head L)) ofr A L C" by (erule valid.induct, auto simp: chain_def sign_def anchor_def) -lemma insertion_resilience: "ALL L. L:valid A n B --> Suc i < len L ---> ins (L,Suc i,M) ~:valid A n B" +lemma insertion_resilience: "\L. L \ valid A n B \ Suc i < len L +\ ins (L,Suc i,M) \ valid A n B" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l', simp) -apply (ind_cases "\x,M,l'\:valid A n B" for x l', clarsimp) -apply (ind_cases "\head l',l'\:valid A n B" for l', simp, simp) +apply (ind_cases "\x,l'\ \ valid A n B" for x l', simp) +apply (ind_cases "\x,M,l'\ \ valid A n B" for x l', clarsimp) +apply (ind_cases "\head l',l'\ \ valid A n B" for l', simp, simp) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,ins(l',Suc na,M)\:valid A n B" for x l' na) +apply (ind_cases "\x,ins(l',Suc na,M)\ \ valid A n B" for x l' na) apply (frule len_not_empty, clarsimp) by (drule_tac x=l' in spec, clarsimp) subsection\truncation resilience: only shop i can truncate at offer i\ -lemma truncation_resilience: "ALL L. L:valid A n B --> Suc i < len L ---> cons M (trunc (L,Suc i)):valid A n B --> shop M = shop (ith (L,i))" +lemma truncation_resilience: "\L. L \ valid A n B \ Suc i < len L +\ cons M (trunc (L,Suc i)) \ valid A n B \ shop M = shop (ith (L,i))" apply (induct i) (* i = 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) -apply (ind_cases "\M,l'\:valid A n B" for l') +apply (ind_cases "\M,l'\ \ valid A n B" for l') apply (frule len_not_empty, clarsimp, simp) (* i > 0 *) apply clarify apply (frule len_not_empty, clarsimp) -apply (ind_cases "\x,l'\:valid A n B" for x l') +apply (ind_cases "\x,l'\ \ valid A n B" for x l') apply (frule len_not_empty, clarsimp) by (drule_tac x=l' in spec, clarsimp) @@ -237,37 +237,37 @@ subsection\get components of a message\ -lemma get_ML [dest]: "Says A' B \A,R,I,M,L\:set evs ==> -M:parts (spies evs) & L:parts (spies evs)" +lemma get_ML [dest]: "Says A' B \A,R,I,M,L\ \ set evs \ +M \ parts (spies evs) \ L \ parts (spies evs)" by blast subsection\general properties of p2\ lemma reqm_neq_prom [iff]: -"reqm A r n I B ~= prom B' ofr A' r' I' (cons M L) J C" +"reqm A r n I B \ prom B' ofr A' r' I' (cons M L) J C" by (auto simp: reqm_def prom_def) lemma prom_neq_reqm [iff]: -"prom B' ofr A' r' I' (cons M L) J C ~= reqm A r n I B" +"prom B' ofr A' r' I' (cons M L) J C \ reqm A r n I B" by (auto simp: reqm_def prom_def) -lemma req_neq_pro [iff]: "req A r n I B ~= pro B' ofr A' r' I' (cons M L) J C" +lemma req_neq_pro [iff]: "req A r n I B \ pro B' ofr A' r' I' (cons M L) J C" by (auto simp: req_def pro_def) -lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C ~= req A r n I B" +lemma pro_neq_req [iff]: "pro B' ofr A' r' I' (cons M L) J C \ req A r n I B" by (auto simp: req_def pro_def) -lemma p2_has_no_Gets: "evs:p2 ==> ALL A X. Gets A X ~:set evs" +lemma p2_has_no_Gets: "evs \ p2 \ \A X. Gets A X \ set evs" by (erule p2.induct, auto simp: req_def pro_def) lemma p2_is_Gets_correct [iff]: "Gets_correct p2" by (auto simp: Gets_correct_def dest: p2_has_no_Gets) lemma p2_is_one_step [iff]: "one_step p2" -by (unfold one_step_def, clarify, ind_cases "ev#evs:p2" for ev evs, auto) +by (unfold one_step_def, clarify, ind_cases "ev#evs \ p2" for ev evs, auto) -lemma p2_has_only_Says' [rule_format]: "evs:p2 ==> -ev:set evs --> (EX A B X. ev=Says A B X)" +lemma p2_has_only_Says' [rule_format]: "evs \ p2 \ +ev \ set evs \ (\A B X. ev=Says A B X)" by (erule p2.induct, auto simp: req_def pro_def) lemma p2_has_only_Says [iff]: "has_only_Says p2" @@ -283,8 +283,8 @@ subsection\private keys are safe\ lemma priK_parts_Friend_imp_bad [rule_format,dest]: - "[| evs:p2; Friend B ~= A |] - ==> (Key (priK A):parts (knows (Friend B) evs)) --> (A:bad)" + "[| evs \ p2; Friend B \ A |] + ==> (Key (priK A) \ parts (knows (Friend B) evs)) \ (A \ bad)" apply (erule p2.induct) apply (simp_all add: initState.simps knows.simps pro_def prom_def req_def reqm_def anchor_def chain_def sign_def) @@ -294,13 +294,13 @@ done lemma priK_analz_Friend_imp_bad [rule_format,dest]: - "[| evs:p2; Friend B ~= A |] -==> (Key (priK A):analz (knows (Friend B) evs)) --> (A:bad)" + "[| evs \ p2; Friend B \ A |] +==> (Key (priK A) \ analz (knows (Friend B) evs)) \ (A \ bad)" by auto lemma priK_notin_knows_max_Friend: - "[| evs:p2; A ~:bad; A ~= Friend C |] - ==> Key (priK A) ~:analz (knows_max (Friend C) evs)" + "[| evs \ p2; A \ bad; A \ Friend C |] + ==> Key (priK A) \ analz (knows_max (Friend C) evs)" apply (rule not_parts_not_analz, simp add: knows_max_def, safe) apply (drule_tac H="spies' evs" in parts_sub) apply (rule_tac p=p2 in knows_max'_sub_spies', simp+) @@ -309,78 +309,78 @@ subsection\general guardedness properties\ -lemma agl_guard [intro]: "I:agl ==> I:guard n Ks" +lemma agl_guard [intro]: "I \ agl \ I \ guard n Ks" by (erule agl.induct, auto) -lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\:set evs; -Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" +lemma Says_to_knows_max'_guard: "[| Says A' C \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" by (auto dest: Says_to_knows_max') -lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\:set evs; -Guard n Ks (knows_max' C evs) |] ==> L:guard n Ks" +lemma Says_from_knows_max'_guard: "[| Says C A' \A'',r,I,L\ \ set evs; +Guard n Ks (knows_max' C evs) |] ==> L \ guard n Ks" by (auto dest: Says_from_knows_max') -lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\:set evs; -Nonce n ~:used evs |] ==> L:guard n Ks" +lemma Says_Nonce_not_used_guard: "[| Says A' B \A'',r,I,L\ \ set evs; +Nonce n \ used evs |] ==> L \ guard n Ks" by (drule not_used_not_parts, auto) subsection\guardedness of messages\ -lemma chain_guard [iff]: "chain B ofr A L C:guard n {priK A}" +lemma chain_guard [iff]: "chain B ofr A L C \ guard n {priK A}" by (case_tac "ofr=n", auto simp: chain_def sign_def) -lemma chain_guard_Nonce_neq [intro]: "n ~= ofr -==> chain B ofr A' L C:guard n {priK A}" +lemma chain_guard_Nonce_neq [intro]: "n \ ofr +\ chain B ofr A' L C \ guard n {priK A}" by (auto simp: chain_def sign_def) -lemma anchor_guard [iff]: "anchor A n' B:guard n {priK A}" +lemma anchor_guard [iff]: "anchor A n' B \ guard n {priK A}" by (case_tac "n'=n", auto simp: anchor_def) -lemma anchor_guard_Nonce_neq [intro]: "n ~= n' -==> anchor A' n' B:guard n {priK A}" +lemma anchor_guard_Nonce_neq [intro]: "n \ n' +\ anchor A' n' B \ guard n {priK A}" by (auto simp: anchor_def) -lemma reqm_guard [intro]: "I:agl ==> reqm A r n' I B:guard n {priK A}" +lemma reqm_guard [intro]: "I \ agl \ reqm A r n' I B \ guard n {priK A}" by (case_tac "n'=n", auto simp: reqm_def) -lemma reqm_guard_Nonce_neq [intro]: "[| n ~= n'; I:agl |] -==> reqm A' r n' I B:guard n {priK A}" +lemma reqm_guard_Nonce_neq [intro]: "[| n \ n'; I \ agl |] +==> reqm A' r n' I B \ guard n {priK A}" by (auto simp: reqm_def) -lemma prom_guard [intro]: "[| I:agl; J:agl; L:guard n {priK A} |] -==> prom B ofr A r I L J C:guard n {priK A}" +lemma prom_guard [intro]: "[| I \ agl; J \ agl; L \ guard n {priK A} |] +==> prom B ofr A r I L J C \ guard n {priK A}" by (auto simp: prom_def) -lemma prom_guard_Nonce_neq [intro]: "[| n ~= ofr; I:agl; J:agl; -L:guard n {priK A} |] ==> prom B ofr A' r I L J C:guard n {priK A}" +lemma prom_guard_Nonce_neq [intro]: "[| n \ ofr; I \ agl; J \ agl; +L \ guard n {priK A} |] ==> prom B ofr A' r I L J C \ guard n {priK A}" by (auto simp: prom_def) subsection\Nonce uniqueness\ -lemma uniq_Nonce_in_chain [dest]: "Nonce k:parts {chain B ofr A L C} ==> k=ofr" +lemma uniq_Nonce_in_chain [dest]: "Nonce k \ parts {chain B ofr A L C} \ k=ofr" by (auto simp: chain_def sign_def) -lemma uniq_Nonce_in_anchor [dest]: "Nonce k:parts {anchor A n B} ==> k=n" +lemma uniq_Nonce_in_anchor [dest]: "Nonce k \ parts {anchor A n B} \ k=n" by (auto simp: anchor_def chain_def sign_def) -lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k:parts {reqm A r n I B}; -I:agl |] ==> k=n" +lemma uniq_Nonce_in_reqm [dest]: "[| Nonce k \ parts {reqm A r n I B}; +I \ agl |] ==> k=n" by (auto simp: reqm_def dest: no_Nonce_in_agl) -lemma uniq_Nonce_in_prom [dest]: "[| Nonce k:parts {prom B ofr A r I L J C}; -I:agl; J:agl; Nonce k ~:parts {L} |] ==> k=ofr" +lemma uniq_Nonce_in_prom [dest]: "[| Nonce k \ parts {prom B ofr A r I L J C}; +I \ agl; J \ agl; Nonce k \ parts {L} |] ==> k=ofr" by (auto simp: prom_def dest: no_Nonce_in_agl no_Nonce_in_appdel) subsection\requests are guarded\ -lemma req_imp_Guard [rule_format]: "[| evs:p2; A ~:bad |] ==> -req A r n I B:set evs --> Guard n {priK A} (spies evs)" +lemma req_imp_Guard [rule_format]: "[| evs \ p2; A \ bad |] ==> +req A r n I B \ set evs \ Guard n {priK A} (spies evs)" apply (erule p2.induct, simp) apply (simp add: req_def knows.simps, safe) apply (erule in_synth_Guard, erule Guard_analz, simp) by (auto simp: req_def pro_def dest: Says_imp_knows_Spy) -lemma req_imp_Guard_Friend: "[| evs:p2; A ~:bad; req A r n I B:set evs |] +lemma req_imp_Guard_Friend: "[| evs \ p2; A \ bad; req A r n I B \ set evs |] ==> Guard n {priK A} (knows_max (Friend C) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) @@ -391,8 +391,8 @@ subsection\propositions are guarded\ -lemma pro_imp_Guard [rule_format]: "[| evs:p2; B ~:bad; A ~:bad |] ==> -pro B ofr A r I (cons M L) J C:set evs --> Guard ofr {priK A} (spies evs)" +lemma pro_imp_Guard [rule_format]: "[| evs \ p2; B \ bad; A \ bad |] ==> +pro B ofr A r I (cons M L) J C \ set evs \ Guard ofr {priK A} (spies evs)" apply (erule p2.induct) (* +3 subgoals *) (* Nil *) apply simp @@ -428,8 +428,8 @@ apply (simp add: pro_def) by (blast dest: Says_imp_knows_Spy) -lemma pro_imp_Guard_Friend: "[| evs:p2; B ~:bad; A ~:bad; -pro B ofr A r I (cons M L) J C:set evs |] +lemma pro_imp_Guard_Friend: "[| evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs |] ==> Guard ofr {priK A} (knows_max (Friend D) evs)" apply (rule Guard_knows_max') apply (rule_tac H="spies evs" in Guard_mono) @@ -441,23 +441,23 @@ subsection\data confidentiality: no one other than the originator can decrypt the offers\ -lemma Nonce_req_notin_spies: "[| evs:p2; req A r n I B:set evs; A ~:bad |] -==> Nonce n ~:analz (spies evs)" +lemma Nonce_req_notin_spies: "[| evs \ p2; req A r n I B \ set evs; A \ bad |] +==> Nonce n \ analz (spies evs)" by (frule req_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_req_notin_knows_max_Friend: "[| evs:p2; req A r n I B:set evs; -A ~:bad; A ~= Friend C |] ==> Nonce n ~:analz (knows_max (Friend C) evs)" +lemma Nonce_req_notin_knows_max_Friend: "[| evs \ p2; req A r n I B \ set evs; +A \ bad; A \ Friend C |] ==> Nonce n \ analz (knows_max (Friend C) evs)" apply (clarify, frule_tac C=C in req_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) -lemma Nonce_pro_notin_spies: "[| evs:p2; B ~:bad; A ~:bad; -pro B ofr A r I (cons M L) J C:set evs |] ==> Nonce ofr ~:analz (spies evs)" +lemma Nonce_pro_notin_spies: "[| evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs |] ==> Nonce ofr \ analz (spies evs)" by (frule pro_imp_Guard, simp+, erule Guard_Nonce_analz, simp+) -lemma Nonce_pro_notin_knows_max_Friend: "[| evs:p2; B ~:bad; A ~:bad; -A ~= Friend D; pro B ofr A r I (cons M L) J C:set evs |] -==> Nonce ofr ~:analz (knows_max (Friend D) evs)" +lemma Nonce_pro_notin_knows_max_Friend: "[| evs \ p2; B \ bad; A \ bad; +A \ Friend D; pro B ofr A r I (cons M L) J C \ set evs |] +==> Nonce ofr \ analz (knows_max (Friend D) evs)" apply (clarify, frule_tac A=A in pro_imp_Guard_Friend, simp+) apply (simp add: knows_max_def, drule Guard_invKey_keyset, simp+) by (drule priK_notin_knows_max_Friend, auto simp: knows_max_def) @@ -465,71 +465,71 @@ subsection\forward privacy: only the originator can know the identity of the shops\ -lemma forward_privacy_Spy: "[| evs:p2; B ~:bad; A ~:bad; -pro B ofr A r I (cons M L) J C:set evs |] -==> sign B (Nonce ofr) ~:analz (spies evs)" +lemma forward_privacy_Spy: "[| evs \ p2; B \ bad; A \ bad; +pro B ofr A r I (cons M L) J C \ set evs |] +==> sign B (Nonce ofr) \ analz (spies evs)" by (auto simp:sign_def dest: Nonce_pro_notin_spies) -lemma forward_privacy_Friend: "[| evs:p2; B ~:bad; A ~:bad; A ~= Friend D; -pro B ofr A r I (cons M L) J C:set evs |] -==> sign B (Nonce ofr) ~:analz (knows_max (Friend D) evs)" +lemma forward_privacy_Friend: "[| evs \ p2; B \ bad; A \ bad; A \ Friend D; +pro B ofr A r I (cons M L) J C \ set evs |] +==> sign B (Nonce ofr) \ analz (knows_max (Friend D) evs)" by (auto simp:sign_def dest:Nonce_pro_notin_knows_max_Friend ) subsection\non repudiability: an offer signed by B has been sent by B\ -lemma Crypt_reqm: "[| Crypt (priK A) X:parts {reqm A' r n I B}; I:agl |] ==> A=A'" +lemma Crypt_reqm: "[| Crypt (priK A) X \ parts {reqm A' r n I B}; I \ agl |] ==> A=A'" by (auto simp: reqm_def anchor_def chain_def sign_def dest: no_Crypt_in_agl) -lemma Crypt_prom: "[| Crypt (priK A) X:parts {prom B ofr A' r I L J C}; -I:agl; J:agl |] ==> A=B | Crypt (priK A) X:parts {L}" +lemma Crypt_prom: "[| Crypt (priK A) X \ parts {prom B ofr A' r I L J C}; +I \ agl; J \ agl |] ==> A=B | Crypt (priK A) X \ parts {L}" apply (simp add: prom_def anchor_def chain_def sign_def) by (blast dest: no_Crypt_in_agl no_Crypt_in_appdel) -lemma Crypt_safeness: "[| evs:p2; A ~:bad |] ==> Crypt (priK A) X:parts (spies evs) ---> (EX B Y. Says A B Y:set evs & Crypt (priK A) X:parts {Y})" +lemma Crypt_safeness: "[| evs \ p2; A \ bad |] ==> Crypt (priK A) X \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs & Crypt (priK A) X \ parts {Y})" apply (erule p2.induct) (* Nil *) apply simp (* Fake *) apply clarsimp -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (erule disjE) apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) (* Request *) apply (simp add: req_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (erule disjE) apply (frule Crypt_reqm, simp, clarify) apply (rule_tac x=B in exI, rule_tac x="reqm A r n I B" in exI, simp, blast) (* Propose *) apply (simp add: pro_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) X:G" in parts_insert_substD, simp) +apply (drule_tac P="\G. Crypt (priK A) X \ G" in parts_insert_substD, simp) apply (rotate_tac -1, erule disjE) apply (frule Crypt_prom, simp, simp) apply (rotate_tac -1, erule disjE) apply (rule_tac x=C in exI) apply (rule_tac x="prom B ofr Aa r I (cons M L) J C" in exI, blast) -apply (subgoal_tac "cons M L:parts (spies evsp)") +apply (subgoal_tac "cons M L \ parts (spies evsp)") apply (drule_tac G="{cons M L}" and H="spies evsp" in parts_trans, blast, blast) apply (drule Says_imp_spies, rotate_tac -1, drule parts.Inj) apply (drule parts.Snd, drule parts.Snd, drule parts.Snd) by auto -lemma Crypt_Hash_imp_sign: "[| evs:p2; A ~:bad |] ==> -Crypt (priK A) (Hash X):parts (spies evs) ---> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" +lemma Crypt_Hash_imp_sign: "[| evs \ p2; A \ bad |] ==> +Crypt (priK A) (Hash X) \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (erule p2.induct) (* Nil *) apply simp (* Fake *) apply clarsimp -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (erule disjE) apply (drule_tac K="priK A" in Crypt_synth, simp+, blast, blast) (* Request *) apply (simp add: req_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (erule disjE) apply (frule Crypt_reqm, simp+) @@ -538,7 +538,7 @@ apply (simp add: chain_def sign_def, blast) (* Propose *) apply (simp add: pro_def, clarify) -apply (drule_tac P="%G. Crypt (priK A) (Hash X):G" in parts_insert_substD) +apply (drule_tac P="\G. Crypt (priK A) (Hash X) \ G" in parts_insert_substD) apply simp apply (rotate_tac -1, erule disjE) apply (simp add: prom_def sign_def no_Crypt_in_agl no_Crypt_in_appdel) @@ -552,8 +552,8 @@ apply (blast del: MPair_parts)+ done -lemma sign_safeness: "[| evs:p2; A ~:bad |] ==> sign A X:parts (spies evs) ---> (EX B Y. Says A B Y:set evs & sign A X:parts {Y})" +lemma sign_safeness: "[| evs \ p2; A \ bad |] ==> sign A X \ parts (spies evs) +\ (\B Y. Says A B Y \ set evs \ sign A X \ parts {Y})" apply (clarify, simp add: sign_def, frule parts.Snd) apply (blast dest: Crypt_Hash_imp_sign [unfolded sign_def]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Guard/Proto.thy --- a/src/HOL/Auth/Guard/Proto.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Guard/Proto.thy Thu Feb 15 13:04:36 2018 +0100 @@ -18,8 +18,8 @@ type_synonym proto = "rule set" definition wdef :: "proto => bool" where -"wdef p == ALL R k. R:p --> Number k:parts {msg' R} ---> Number k:parts (msg`(fst R))" +"wdef p \ \R k. R \ p \ Number k \ parts {msg' R} +\ Number k \ parts (msg`(fst R))" subsection\substitutions\ @@ -36,89 +36,89 @@ | "apm s (Key K) = Key (key s K)" | "apm s (Hash X) = Hash (apm s X)" | "apm s (Crypt K X) = ( -if (EX A. K = pubK A) then Crypt (pubK (agent s (agt K))) (apm s X) -else if (EX A. K = priK A) then Crypt (priK (agent s (agt K))) (apm s X) +if (\A. K = pubK A) then Crypt (pubK (agent s (agt K))) (apm s X) +else if (\A. K = priK A) then Crypt (priK (agent s (agt K))) (apm s X) else Crypt (key s K) (apm s X))" | "apm s \X,Y\ = \apm s X, apm s Y\" -lemma apm_parts: "X:parts {Y} ==> apm s X:parts {apm s Y}" +lemma apm_parts: "X \ parts {Y} \ apm s X \ parts {apm s Y}" apply (erule parts.induct, simp_all, blast) apply (erule parts.Fst) apply (erule parts.Snd) by (erule parts.Body)+ -lemma Nonce_apm [rule_format]: "Nonce n:parts {apm s X} ==> -(ALL k. Number k:parts {X} --> Nonce n ~:parts {nb s k}) --> -(EX k. Nonce k:parts {X} & nonce s k = n)" +lemma Nonce_apm [rule_format]: "Nonce n \ parts {apm s X} \ +(\k. Number k \ parts {X} \ Nonce n \ parts {nb s k}) \ +(\k. Nonce k \ parts {X} \ nonce s k = n)" by (induct X, simp_all, blast) -lemma wdef_Nonce: "[| Nonce n:parts {apm s X}; R:p; msg' R = X; wdef p; -Nonce n ~:parts (apm s `(msg `(fst R))) |] ==> -(EX k. Nonce k:parts {X} & nonce s k = n)" +lemma wdef_Nonce: "[| Nonce n \ parts {apm s X}; R \ p; msg' R = X; wdef p; +Nonce n \ parts (apm s `(msg `(fst R))) |] ==> +(\k. Nonce k \ parts {X} \ nonce s k = n)" apply (erule Nonce_apm, unfold wdef_def) apply (drule_tac x=R in spec, drule_tac x=k in spec, clarsimp) apply (drule_tac x=x in bspec, simp) apply (drule_tac Y="msg x" and s=s in apm_parts, simp) by (blast dest: parts_parts) -primrec ap :: "subs => event => event" where +primrec ap :: "subs \ event \ event" where "ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)" | "ap s (Gets A X) = Gets (agent s A) (apm s X)" | "ap s (Notes A X) = Notes (agent s A) (apm s X)" abbreviation - ap' :: "subs => rule => event" where - "ap' s R == ap s (snd R)" + ap' :: "subs \ rule \ event" where + "ap' s R \ ap s (snd R)" abbreviation - apm' :: "subs => rule => msg" where - "apm' s R == apm s (msg' R)" + apm' :: "subs \ rule \ msg" where + "apm' s R \ apm s (msg' R)" abbreviation - priK' :: "subs => agent => key" where - "priK' s A == priK (agent s A)" + priK' :: "subs \ agent \ key" where + "priK' s A \ priK (agent s A)" abbreviation - pubK' :: "subs => agent => key" where - "pubK' s A == pubK (agent s A)" + pubK' :: "subs \ agent \ key" where + "pubK' s A \ pubK (agent s A)" subsection\nonces generated by a rule\ -definition newn :: "rule => nat set" where -"newn R == {n. Nonce n:parts {msg (snd R)} & Nonce n ~:parts (msg`(fst R))}" +definition newn :: "rule \ nat set" where +"newn R \ {n. Nonce n \ parts {msg (snd R)} \ Nonce n \ parts (msg`(fst R))}" -lemma newn_parts: "n:newn R ==> Nonce (nonce s n):parts {apm' s R}" +lemma newn_parts: "n \ newn R \ Nonce (nonce s n) \ parts {apm' s R}" by (auto simp: newn_def dest: apm_parts) subsection\traces generated by a protocol\ -definition ok :: "event list => rule => subs => bool" where -"ok evs R s == ((ALL x. x:fst R --> ap s x:set evs) -& (ALL n. n:newn R --> Nonce (nonce s n) ~:used evs))" +definition ok :: "event list \ rule \ subs \ bool" where +"ok evs R s \ ((\x. x \ fst R \ ap s x \ set evs) +\ (\n. n \ newn R \ Nonce (nonce s n) \ used evs))" inductive_set tr :: "proto => event list set" for p :: proto where - Nil [intro]: "[]:tr p" + Nil [intro]: "[] \ tr p" -| Fake [intro]: "[| evsf:tr p; X:synth (analz (spies evsf)) |] - ==> Says Spy B X # evsf:tr p" +| Fake [intro]: "[| evsf \ tr p; X \ synth (analz (spies evsf)) |] + ==> Says Spy B X # evsf \ tr p" -| Proto [intro]: "[| evs:tr p; R:p; ok evs R s |] ==> ap' s R # evs:tr p" +| Proto [intro]: "[| evs \ tr p; R \ p; ok evs R s |] ==> ap' s R # evs \ tr p" subsection\general properties\ lemma one_step_tr [iff]: "one_step (tr p)" apply (unfold one_step_def, clarify) -by (ind_cases "ev # evs:tr p" for ev evs, auto) +by (ind_cases "ev # evs \ tr p" for ev evs, auto) definition has_only_Says' :: "proto => bool" where -"has_only_Says' p == ALL R. R:p --> is_Says (snd R)" +"has_only_Says' p \ \R. R \ p \ is_Says (snd R)" -lemma has_only_Says'D: "[| R:p; has_only_Says' p |] -==> (EX A B X. snd R = Says A B X)" +lemma has_only_Says'D: "[| R \ p; has_only_Says' p |] +==> (\A B X. snd R = Says A B X)" by (unfold has_only_Says'_def is_Says_def, blast) lemma has_only_Says_tr [simp]: "has_only_Says' p ==> has_only_Says (tr p)" @@ -129,17 +129,17 @@ by (drule_tac x=a in spec, auto simp: is_Says_def) lemma has_only_Says'_in_trD: "[| has_only_Says' p; list @ ev # evs1 \ tr p |] -==> (EX A B X. ev = Says A B X)" +==> (\A B X. ev = Says A B X)" by (drule has_only_Says_tr, auto) -lemma ok_not_used: "[| Nonce n ~:used evs; ok evs R s; -ALL x. x:fst R --> is_Says x |] ==> Nonce n ~:parts (apm s `(msg `(fst R)))" +lemma ok_not_used: "[| Nonce n \ used evs; ok evs R s; +\x. x \ fst R \ is_Says x |] ==> Nonce n \ parts (apm s `(msg `(fst R)))" apply (unfold ok_def, clarsimp) apply (drule_tac x=x in spec, drule_tac x=x in spec) by (auto simp: is_Says_def dest: Says_imp_spies not_used_not_spied parts_parts) -lemma ok_is_Says: "[| evs' @ ev # evs:tr p; ok evs R s; has_only_Says' p; -R:p; x:fst R |] ==> is_Says x" +lemma ok_is_Says: "[| evs' @ ev # evs \ tr p; ok evs R s; has_only_Says' p; +R \ p; x \ fst R |] ==> is_Says x" apply (unfold ok_def is_Says_def, clarify) apply (drule_tac x=x in spec, simp) apply (subgoal_tac "one_step (tr p)") @@ -149,42 +149,42 @@ subsection\types\ -type_synonym keyfun = "rule => subs => nat => event list => key set" +type_synonym keyfun = "rule \ subs \ nat \ event list \ key set" -type_synonym secfun = "rule => nat => subs => key set => msg" +type_synonym secfun = "rule \ nat \ subs \ key set \ msg" subsection\introduction of a fresh guarded nonce\ -definition fresh :: "proto => rule => subs => nat => key set => event list -=> bool" where -"fresh p R s n Ks evs == (EX evs1 evs2. evs = evs2 @ ap' s R # evs1 -& Nonce n ~:used evs1 & R:p & ok evs1 R s & Nonce n:parts {apm' s R} -& apm' s R:guard n Ks)" +definition fresh :: "proto \ rule \ subs \ nat \ key set \ event list +\ bool" where +"fresh p R s n Ks evs \ (\evs1 evs2. evs = evs2 @ ap' s R # evs1 +\ Nonce n \ used evs1 \ R \ p \ ok evs1 R s \ Nonce n \ parts {apm' s R} +\ apm' s R \ guard n Ks)" -lemma freshD: "fresh p R s n Ks evs ==> (EX evs1 evs2. -evs = evs2 @ ap' s R # evs1 & Nonce n ~:used evs1 & R:p & ok evs1 R s -& Nonce n:parts {apm' s R} & apm' s R:guard n Ks)" +lemma freshD: "fresh p R s n Ks evs \ (\evs1 evs2. +evs = evs2 @ ap' s R # evs1 \ Nonce n \ used evs1 \ R \ p \ ok evs1 R s +\ Nonce n \ parts {apm' s R} \ apm' s R \ guard n Ks)" by (unfold fresh_def, blast) -lemma freshI [intro]: "[| Nonce n ~:used evs1; R:p; Nonce n:parts {apm' s R}; -ok evs1 R s; apm' s R:guard n Ks |] +lemma freshI [intro]: "[| Nonce n \ used evs1; R \ p; Nonce n \ parts {apm' s R}; +ok evs1 R s; apm' s R \ guard n Ks |] ==> fresh p R s n Ks (list @ ap' s R # evs1)" by (unfold fresh_def, blast) -lemma freshI': "[| Nonce n ~:used evs1; (l,r):p; -Nonce n:parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r):guard n Ks |] +lemma freshI': "[| Nonce n \ used evs1; (l,r) \ p; +Nonce n \ parts {apm s (msg r)}; ok evs1 (l,r) s; apm s (msg r) \ guard n Ks |] ==> fresh p (l,r) s n Ks (evs2 @ ap s r # evs1)" by (drule freshI, simp+) lemma fresh_used: "[| fresh p R' s' n Ks evs; has_only_Says' p |] -==> Nonce n:used evs" +==> Nonce n \ used evs" apply (unfold fresh_def, clarify) apply (drule has_only_Says'D) by (auto intro: parts_used_app) -lemma fresh_newn: "[| evs' @ ap' s R # evs:tr p; wdef p; has_only_Says' p; -Nonce n ~:used evs; R:p; ok evs R s; Nonce n:parts {apm' s R} |] -==> EX k. k:newn R & nonce s k = n" +lemma fresh_newn: "[| evs' @ ap' s R # evs \ tr p; wdef p; has_only_Says' p; +Nonce n \ used evs; R \ p; ok evs R s; Nonce n \ parts {apm' s R} |] +==> \k. k \ newn R \ nonce s k = n" apply (drule wdef_Nonce, simp+) apply (frule ok_not_used, simp+) apply (clarify, erule ok_is_Says, simp+) @@ -193,22 +193,22 @@ apply (drule ok_not_used, simp+) by (clarify, erule ok_is_Says, simp_all) -lemma fresh_rule: "[| evs' @ ev # evs:tr p; wdef p; Nonce n ~:used evs; -Nonce n:parts {msg ev} |] ==> EX R s. R:p & ap' s R = ev" -apply (drule trunc, simp, ind_cases "ev # evs:tr p", simp) +lemma fresh_rule: "[| evs' @ ev # evs \ tr p; wdef p; Nonce n \ used evs; +Nonce n \ parts {msg ev} |] ==> \R s. R \ p \ ap' s R = ev" +apply (drule trunc, simp, ind_cases "ev # evs \ tr p", simp) by (drule_tac x=X in in_sub, drule parts_sub, simp, simp, blast+) -lemma fresh_ruleD: "[| fresh p R' s' n Ks evs; keys R' s' n evs <= Ks; wdef p; -has_only_Says' p; evs:tr p; ALL R k s. nonce s k = n --> Nonce n:used evs --> -R:p --> k:newn R --> Nonce n:parts {apm' s R} --> apm' s R:guard n Ks --> -apm' s R:parts (spies evs) --> keys R s n evs <= Ks --> P |] ==> P" +lemma fresh_ruleD: "[| fresh p R' s' n Ks evs; keys R' s' n evs \ Ks; wdef p; +has_only_Says' p; evs \ tr p; \R k s. nonce s k = n \ Nonce n \ used evs \ +R \ p \ k \ newn R \ Nonce n \ parts {apm' s R} \ apm' s R \ guard n Ks \ +apm' s R \ parts (spies evs) \ keys R s n evs \ Ks \ P |] ==> P" apply (frule fresh_used, simp) apply (unfold fresh_def, clarify) apply (drule_tac x=R' in spec) apply (drule fresh_newn, simp+, clarify) apply (drule_tac x=k in spec) apply (drule_tac x=s' in spec) -apply (subgoal_tac "apm' s' R':parts (spies (evs2 @ ap' s' R' # evs1))") +apply (subgoal_tac "apm' s' R' \ parts (spies (evs2 @ ap' s' R' # evs1))") apply (case_tac R', drule has_only_Says'D, simp, clarsimp) apply (case_tac R', drule has_only_Says'D, simp, clarsimp) apply (rule_tac Y="apm s' X" in parts_parts, blast) @@ -216,50 +216,50 @@ subsection\safe keys\ -definition safe :: "key set => msg set => bool" where -"safe Ks G == ALL K. K:Ks --> Key K ~:analz G" +definition safe :: "key set \ msg set \ bool" where +"safe Ks G \ \K. K \ Ks \ Key K \ analz G" -lemma safeD [dest]: "[| safe Ks G; K:Ks |] ==> Key K ~:analz G" +lemma safeD [dest]: "[| safe Ks G; K \ Ks |] ==> Key K \ analz G" by (unfold safe_def, blast) lemma safe_insert: "safe Ks (insert X G) ==> safe Ks G" by (unfold safe_def, blast) -lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n ~:analz G" +lemma Guard_safe: "[| Guard n Ks G; safe Ks G |] ==> Nonce n \ analz G" by (blast dest: Guard_invKey) subsection\guardedness preservation\ -definition preserv :: "proto => keyfun => nat => key set => bool" where -"preserv p keys n Ks == (ALL evs R' s' R s. evs:tr p --> -Guard n Ks (spies evs) --> safe Ks (spies evs) --> fresh p R' s' n Ks evs --> -keys R' s' n evs <= Ks --> R:p --> ok evs R s --> apm' s R:guard n Ks)" +definition preserv :: "proto \ keyfun \ nat \ key set \ bool" where +"preserv p keys n Ks \ (\evs R' s' R s. evs \ tr p \ +Guard n Ks (spies evs) \ safe Ks (spies evs) \ fresh p R' s' n Ks evs \ +keys R' s' n evs \ Ks \ R \ p \ ok evs R s \ apm' s R \ guard n Ks)" -lemma preservD: "[| preserv p keys n Ks; evs:tr p; Guard n Ks (spies evs); -safe Ks (spies evs); fresh p R' s' n Ks evs; R:p; ok evs R s; -keys R' s' n evs <= Ks |] ==> apm' s R:guard n Ks" +lemma preservD: "[| preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); +safe Ks (spies evs); fresh p R' s' n Ks evs; R \ p; ok evs R s; +keys R' s' n evs \ Ks |] ==> apm' s R \ guard n Ks" by (unfold preserv_def, blast) -lemma preservD': "[| preserv p keys n Ks; evs:tr p; Guard n Ks (spies evs); -safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X):p; -ok evs (l,Says A B X) s; keys R' s' n evs <= Ks |] ==> apm s X:guard n Ks" +lemma preservD': "[| preserv p keys n Ks; evs \ tr p; Guard n Ks (spies evs); +safe Ks (spies evs); fresh p R' s' n Ks evs; (l,Says A B X) \ p; +ok evs (l,Says A B X) s; keys R' s' n evs \ Ks |] ==> apm s X \ guard n Ks" by (drule preservD, simp+) subsection\monotonic keyfun\ definition monoton :: "proto => keyfun => bool" where -"monoton p keys == ALL R' s' n ev evs. ev # evs:tr p --> -keys R' s' n evs <= keys R' s' n (ev # evs)" +"monoton p keys \ \R' s' n ev evs. ev # evs \ tr p \ +keys R' s' n evs \ keys R' s' n (ev # evs)" -lemma monotonD [dest]: "[| keys R' s' n (ev # evs) <= Ks; monoton p keys; -ev # evs:tr p |] ==> keys R' s' n evs <= Ks" +lemma monotonD [dest]: "[| keys R' s' n (ev # evs) \ Ks; monoton p keys; +ev # evs \ tr p |] ==> keys R' s' n evs \ Ks" by (unfold monoton_def, blast) subsection\guardedness theorem\ -lemma Guard_tr [rule_format]: "[| evs:tr p; has_only_Says' p; +lemma Guard_tr [rule_format]: "[| evs \ tr p; has_only_Says' p; preserv p keys n Ks; monoton p keys; Guard n Ks (initState Spy) |] ==> -safe Ks (spies evs) --> fresh p R' s' n Ks evs --> keys R' s' n evs <= Ks --> +safe Ks (spies evs) \ fresh p R' s' n Ks evs \ keys R' s' n evs \ Ks \ Guard n Ks (spies evs)" apply (erule tr.induct) (* Nil *) @@ -297,59 +297,59 @@ subsection\useful properties for guardedness\ -lemma newn_neq_used: "[| Nonce n:used evs; ok evs R s; k:newn R |] -==> n ~= nonce s k" +lemma newn_neq_used: "[| Nonce n \ used evs; ok evs R s; k \ newn R |] +==> n \ nonce s k" by (auto simp: ok_def) -lemma ok_Guard: "[| ok evs R s; Guard n Ks (spies evs); x:fst R; is_Says x |] -==> apm s (msg x):parts (spies evs) & apm s (msg x):guard n Ks" +lemma ok_Guard: "[| ok evs R s; Guard n Ks (spies evs); x \ fst R; is_Says x |] +==> apm s (msg x) \ parts (spies evs) \ apm s (msg x) \ guard n Ks" apply (unfold ok_def is_Says_def, clarify) apply (drule_tac x="Says A B X" in spec, simp) by (drule Says_imp_spies, auto intro: parts_parts) -lemma ok_parts_not_new: "[| Y:parts (spies evs); Nonce (nonce s n):parts {Y}; -ok evs R s |] ==> n ~:newn R" +lemma ok_parts_not_new: "[| Y \ parts (spies evs); Nonce (nonce s n) \ parts {Y}; +ok evs R s |] ==> n \ newn R" by (auto simp: ok_def dest: not_used_not_spied parts_parts) subsection\unicity\ -definition uniq :: "proto => secfun => bool" where -"uniq p secret == ALL evs R R' n n' Ks s s'. R:p --> R':p --> -n:newn R --> n':newn R' --> nonce s n = nonce s' n' --> -Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} --> -apm' s R:guard (nonce s n) Ks --> apm' s' R':guard (nonce s n) Ks --> -evs:tr p --> Nonce (nonce s n) ~:analz (spies evs) --> -secret R n s Ks:parts (spies evs) --> secret R' n' s' Ks:parts (spies evs) --> +definition uniq :: "proto \ secfun \ bool" where +"uniq p secret \ \evs R R' n n' Ks s s'. R \ p \ R' \ p \ +n \ newn R \ n' \ newn R' \ nonce s n = nonce s' n' \ +Nonce (nonce s n) \ parts {apm' s R} \ Nonce (nonce s n) \ parts {apm' s' R'} \ +apm' s R \ guard (nonce s n) Ks \ apm' s' R' \ guard (nonce s n) Ks \ +evs \ tr p \ Nonce (nonce s n) \ analz (spies evs) \ +secret R n s Ks \ parts (spies evs) \ secret R' n' s' Ks \ parts (spies evs) \ secret R n s Ks = secret R' n' s' Ks" -lemma uniqD: "[| uniq p secret; evs: tr p; R:p; R':p; n:newn R; n':newn R'; -nonce s n = nonce s' n'; Nonce (nonce s n) ~:analz (spies evs); -Nonce (nonce s n):parts {apm' s R}; Nonce (nonce s n):parts {apm' s' R'}; -secret R n s Ks:parts (spies evs); secret R' n' s' Ks:parts (spies evs); -apm' s R:guard (nonce s n) Ks; apm' s' R':guard (nonce s n) Ks |] ==> +lemma uniqD: "[| uniq p secret; evs \ tr p; R \ p; R' \ p; n \ newn R; n' \ newn R'; +nonce s n = nonce s' n'; Nonce (nonce s n) \ analz (spies evs); +Nonce (nonce s n) \ parts {apm' s R}; Nonce (nonce s n) \ parts {apm' s' R'}; +secret R n s Ks \ parts (spies evs); secret R' n' s' Ks \ parts (spies evs); +apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks |] ==> secret R n s Ks = secret R' n' s' Ks" by (unfold uniq_def, blast) -definition ord :: "proto => (rule => rule => bool) => bool" where -"ord p inff == ALL R R'. R:p --> R':p --> ~ inff R R' --> inff R' R" +definition ord :: "proto \ (rule \ rule \ bool) \ bool" where +"ord p inff \ \R R'. R \ p \ R' \ p \ \ inff R R' \ inff R' R" -lemma ordD: "[| ord p inff; ~ inff R R'; R:p; R':p |] ==> inff R' R" +lemma ordD: "[| ord p inff; \ inff R R'; R \ p; R' \ p |] ==> inff R' R" by (unfold ord_def, blast) -definition uniq' :: "proto => (rule => rule => bool) => secfun => bool" where -"uniq' p inff secret == ALL evs R R' n n' Ks s s'. R:p --> R':p --> -inff R R' --> n:newn R --> n':newn R' --> nonce s n = nonce s' n' --> -Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} --> -apm' s R:guard (nonce s n) Ks --> apm' s' R':guard (nonce s n) Ks --> -evs:tr p --> Nonce (nonce s n) ~:analz (spies evs) --> -secret R n s Ks:parts (spies evs) --> secret R' n' s' Ks:parts (spies evs) --> +definition uniq' :: "proto \ (rule \ rule \ bool) \ secfun \ bool" where +"uniq' p inff secret \ \evs R R' n n' Ks s s'. R \ p \ R' \ p \ +inff R R' \ n \ newn R \ n' \ newn R' \ nonce s n = nonce s' n' \ +Nonce (nonce s n) \ parts {apm' s R} \ Nonce (nonce s n) \ parts {apm' s' R'} \ +apm' s R \ guard (nonce s n) Ks \ apm' s' R' \ guard (nonce s n) Ks \ +evs \ tr p \ Nonce (nonce s n) \ analz (spies evs) \ +secret R n s Ks \ parts (spies evs) \ secret R' n' s' Ks \ parts (spies evs) \ secret R n s Ks = secret R' n' s' Ks" -lemma uniq'D: "[| uniq' p inff secret; evs: tr p; inff R R'; R:p; R':p; n:newn R; -n':newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) ~:analz (spies evs); -Nonce (nonce s n):parts {apm' s R}; Nonce (nonce s n):parts {apm' s' R'}; -secret R n s Ks:parts (spies evs); secret R' n' s' Ks:parts (spies evs); -apm' s R:guard (nonce s n) Ks; apm' s' R':guard (nonce s n) Ks |] ==> +lemma uniq'D: "[| uniq' p inff secret; evs \ tr p; inff R R'; R \ p; R' \ p; n \ newn R; +n' \ newn R'; nonce s n = nonce s' n'; Nonce (nonce s n) \ analz (spies evs); +Nonce (nonce s n) \ parts {apm' s R}; Nonce (nonce s n) \ parts {apm' s' R'}; +secret R n s Ks \ parts (spies evs); secret R' n' s' Ks \ parts (spies evs); +apm' s R \ guard (nonce s n) Ks; apm' s' R' \ guard (nonce s n) Ks |] ==> secret R n s Ks = secret R' n' s' Ks" by (unfold uniq'_def, blast) @@ -385,9 +385,9 @@ Says a b (Crypt (pubK b) (Nonce Nb)))" inductive_set ns :: proto where - [iff]: "ns1:ns" -| [iff]: "ns2:ns" -| [iff]: "ns3:ns" + [iff]: "ns1 \ ns" +| [iff]: "ns2 \ ns" +| [iff]: "ns3 \ ns" abbreviation (input) ns3a :: event where diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/KerberosIV.thy Thu Feb 15 13:04:36 2018 +0100 @@ -22,7 +22,7 @@ definition (* authKeys are those contained in an authTicket *) - authKeys :: "event list => key set" where + authKeys :: "event list \ key set" where "authKeys evs = {authK. \A Peer Ta. Says Kas A (Crypt (shrK A) \Key authK, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key authK, Number Ta\) @@ -31,21 +31,21 @@ definition (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) - Issues :: "[agent, agent, msg, event list] => bool" + Issues :: "[agent, agent, msg, event list] \ bool" ("_ Issues _ with _ on _" [50, 0, 0, 50] 50) where "(A Issues B with X on evs) = - (\Y. Says A B Y \ set evs & X \ parts {Y} & - X \ parts (spies (takeWhile (% z. z \ Says A B Y) (rev evs))))" + (\Y. Says A B Y \ set evs \ X \ parts {Y} \ + X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] => event list" ("before _ on _" [0, 50] 50) - where "(before ev on evs) = takeWhile (% z. z ~= ev) (rev evs)" + before :: "[event, event list] \ event list" ("before _ on _" [0, 50] 50) + where "(before ev on evs) = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] => bool" ("Unique _ on _" [0, 50] 50) - where "(Unique ev on evs) = (ev \ set (tl (dropWhile (% z. z \ ev) evs)))" + Unique :: "[event, event list] \ bool" ("Unique _ on _" [0, 50] 50) + where "(Unique ev on evs) = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" consts @@ -79,30 +79,30 @@ abbreviation (*The current time is the length of the trace*) - CT :: "event list=>nat" where + CT :: "event list \ nat" where "CT == length" abbreviation - expiredAK :: "[nat, event list] => bool" where + expiredAK :: "[nat, event list] \ bool" where "expiredAK Ta evs == authKlife + Ta < CT evs" abbreviation - expiredSK :: "[nat, event list] => bool" where + expiredSK :: "[nat, event list] \ bool" where "expiredSK Ts evs == servKlife + Ts < CT evs" abbreviation - expiredA :: "[nat, event list] => bool" where + expiredA :: "[nat, event list] \ bool" where "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] => bool" ("valid _ wrt _" [0, 50] 50) where - "valid T1 wrt T2 == T1 <= replylife + T2" + valid :: "[nat, nat] \ bool" ("valid _ wrt _" [0, 50] 50) where + "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) (* Predicate formalising the association between authKeys and servKeys *) -definition AKcryptSK :: "[key, key, event list] => bool" where +definition AKcryptSK :: "[key, key, event list] \ bool" where "AKcryptSK authK servK evs == \A B Ts. Says Tgs A (Crypt authK @@ -175,7 +175,7 @@ \ set evs4; \ expiredAK Ta evs4; \ expiredA T2 evs4; - servKlife + (CT evs4) <= authKlife + Ta + servKlife + (CT evs4) \ authKlife + Ta \ \ Says Tgs A (Crypt authK \Key servK, Agent B, Number (CT evs4), @@ -267,7 +267,7 @@ done lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = - (if A:bad then insert X (spies evs) else spies evs)" + (if A\bad then insert X (spies evs) else spies evs)" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] a, auto) @@ -282,7 +282,7 @@ lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] -lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" +lemma spies_takeWhile: "spies (takeWhile P evs) \ spies evs" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -341,7 +341,7 @@ lemma Oops_range_spies1: "\ Says Kas A (Crypt KeyA \Key authK, Peer, Ta, authTicket\) \ set evs ; - evs \ kerbIV \ \ authK \ range shrK & authK \ symKeys" + evs \ kerbIV \ \ authK \ range shrK \ authK \ symKeys" apply (erule rev_mp) apply (erule kerbIV.induct, auto) done @@ -355,7 +355,7 @@ lemma Oops_range_spies2: "\ Says Tgs A (Crypt authK \Key servK, Agent B, Ts, servTicket\) \ set evs ; - evs \ kerbIV \ \ servK \ range shrK & servK \ symKeys" + evs \ kerbIV \ \ servK \ range shrK \ servK \ symKeys" apply (erule rev_mp) apply (erule kerbIV.induct, auto) done @@ -379,7 +379,7 @@ by auto lemma Spy_see_shrK_D [dest!]: - "\ Key (shrK A) \ parts (spies evs); evs \ kerbIV \ \ A:bad" + "\ Key (shrK A) \ parts (spies evs); evs \ kerbIV \ \ A\bad" by (blast dest: Spy_see_shrK) lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] @@ -444,7 +444,7 @@ done lemma used_takeWhile_used [rule_format]: - "x : used (takeWhile P X) --> x : used X" + "x \ used (takeWhile P X) \ x \ used X" apply (induct_tac "X") apply simp apply (rename_tac a b) @@ -469,12 +469,12 @@ "\ Says Kas A (Crypt K \Key authK, Agent Peer, Number Ta, authTicket\) \ set evs; evs \ kerbIV \ \ - K = shrK A & Peer = Tgs & - authK \ range shrK & authK \ authKeys evs & authK \ symKeys & - authTicket = (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\) & + K = shrK A \ Peer = Tgs \ + authK \ range shrK \ authK \ authKeys evs \ authK \ symKeys \ + authTicket = (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\) \ Key authK \ used(before Says Kas A (Crypt K \Key authK, Agent Peer, Number Ta, authTicket\) - on evs) & + on evs) \ Ta = CT (before Says Kas A (Crypt K \Key authK, Agent Peer, Number Ta, authTicket\) on evs)" @@ -540,13 +540,13 @@ "\ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, servTicket\) \ set evs; evs \ kerbIV \ - \ B \ Tgs & - authK \ range shrK & authK \ authKeys evs & authK \ symKeys & - servK \ range shrK & servK \ authKeys evs & servK \ symKeys & - servTicket = (Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\) & + \ B \ Tgs \ + authK \ range shrK \ authK \ authKeys evs \ authK \ symKeys \ + servK \ range shrK \ servK \ authKeys evs \ servK \ symKeys \ + servTicket = (Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\) \ Key servK \ used (before Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, servTicket\) - on evs) & + on evs) \ Ts = CT(before Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, servTicket\) on evs) " @@ -572,7 +572,7 @@ \ parts (spies evs); A \ bad; evs \ kerbIV \ - \ authK \ range shrK & authK \ symKeys & + \ authK \ range shrK \ authK \ symKeys \ authTicket = Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Ta\" apply (erule rev_mp) apply (erule kerbIV.induct) @@ -587,7 +587,7 @@ \ parts (spies evs); Key authK \ analz (spies evs); evs \ kerbIV \ - \ servK \ range shrK & servK \ symKeys & + \ servK \ range shrK \ servK \ symKeys \ (\A. servTicket = Crypt (shrK B) \Agent A, Agent B, Key servK, Ts\)" apply (erule rev_mp) apply (erule rev_mp) @@ -601,7 +601,7 @@ "\ Says Kas' A (Crypt (shrK A) \Key authK, Agent Tgs, Ta, authTicket\) \ set evs; evs \ kerbIV \ - \ authK \ range shrK & authK \ symKeys & + \ authK \ range shrK \ authK \ symKeys \ authTicket = Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Ta\ | authTicket \ analz (spies evs)" @@ -612,7 +612,7 @@ "\ Says Tgs' A (Crypt authK \Key servK, Agent B, Ts, servTicket\) \ set evs; authK \ symKeys; evs \ kerbIV \ - \ servK \ range shrK & + \ servK \ range shrK \ (\A. servTicket = Crypt (shrK B) \Agent A, Agent B, Key servK, Ts\) | servTicket \ analz (spies evs)" @@ -718,7 +718,7 @@ \ \Ta. (Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & servKlife + Ts <= authKlife + Ta)" + \ servKlife + Ts \ authKlife + Ta)" apply (erule rev_mp) apply (erule kerbIV.induct) apply (frule_tac [7] K5_msg_in_parts_spies) @@ -744,7 +744,7 @@ \ \authK Ta. Says Kas A (Crypt(shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & servKlife + Ts <= authKlife + Ta" + \ servKlife + Ts \ authKlife + Ta" by (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2) lemma servTicket_authentic: @@ -755,7 +755,7 @@ Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, + \ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\\) \ set evs" by (blast dest: servTicket_authentic_Tgs K4_imp_K2) @@ -768,14 +768,14 @@ (Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, + \ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\\) \ set evs - & servKlife + Ts <= authKlife + Ta)" + \ servKlife + Ts \ authKlife + Ta)" by (blast dest: servTicket_authentic_Tgs u_K4_imp_K2) lemma u_NotexpiredSK_NotexpiredAK: - "\ \ expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \ + "\ \ expiredSK Ts evs; servKlife + Ts \ authKlife + Ta \ \ \ expiredAK Ta evs" by (metis le_less_trans) @@ -804,7 +804,7 @@ Crypt K' \Key SesKey, Agent B', T', Ticket'\ \ parts (spies evs); Key SesKey \ analz (spies evs); evs \ kerbIV \ - \ K=K' & B=B' & T=T' & Ticket=Ticket'" + \ K=K' \ B=B' \ T=T' \ Ticket=Ticket'" apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -870,7 +870,7 @@ Crypt (shrK B') \Agent A', Agent B', Key SesKey, T'\ \ parts (spies evs); Key SesKey \ analz (spies evs); evs \ kerbIV \ - \ A=A' & B=B' & T=T'" + \ A=A' \ B=B' \ T=T'" apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -947,7 +947,7 @@ \ Key Kc \ analz (spies evs) \ (\K' B' T' Ticket'. \K B T Ticket. Crypt Kc \Key K, Agent B, T, Ticket\ - \ parts (spies evs) \ K=K' & B=B' & T=T' & Ticket=Ticket')" + \ parts (spies evs) \ K=K' \ B=B' \ T=T' \ Ticket=Ticket')" would fail on the K2 and K4 cases. *) @@ -957,7 +957,7 @@ (Crypt Ka \Key authK, Agent Tgs, Ta, X\) \ set evs; Says Kas A' (Crypt Ka' \Key authK, Agent Tgs, Ta', X'\) \ set evs; - evs \ kerbIV \ \ A=A' & Ka=Ka' & Ta=Ta' & X=X'" + evs \ kerbIV \ \ A=A' \ Ka=Ka' \ Ta=Ta' \ X=X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule kerbIV.induct) @@ -973,7 +973,7 @@ (Crypt K \Key servK, Agent B, Ts, X\) \ set evs; Says Tgs A' (Crypt K' \Key servK, Agent B', Ts', X'\) \ set evs; - evs \ kerbIV \ \ A=A' & B=B' & K=K' & Ts=Ts' & X=X'" + evs \ kerbIV \ \ A=A' \ B=B' \ K=K' \ Ts=Ts' \ X=X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule kerbIV.induct) @@ -1020,7 +1020,7 @@ lemma AKcryptSK_Says [simp]: "AKcryptSK authK servK (Says S A X # evs) = - (Tgs = S & + (Tgs = S \ (\B Ts. X = Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\ \) @@ -1123,9 +1123,9 @@ text\We take some pains to express the property as a logical equivalence so that the simplifier can apply it.\ lemma Key_analz_image_Key_lemma: - "P \ (Key K \ analz (Key`KK Un H)) \ (K:KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K\KK | Key K \ analz H) \ - P \ (Key K \ analz (Key`KK Un H)) = (K:KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K\KK | Key K \ analz H)" by (blast intro: analz_mono [THEN subsetD]) @@ -1137,7 +1137,7 @@ done lemma authKeys_are_not_AKcryptSK: - "\ K \ authKeys evs Un range shrK; evs \ kerbIV \ + "\ K \ authKeys evs \ range shrK; evs \ kerbIV \ \ \SK. \ AKcryptSK SK K evs \ K \ symKeys" apply (simp add: authKeys_def AKcryptSK_def) apply (blast dest: Says_Kas_message_form Says_Tgs_message_form) @@ -1170,9 +1170,9 @@ [simplified by LCP]\ lemma Key_analz_image_Key [rule_format (no_asm)]: "evs \ kerbIV \ - (\SK KK. SK \ symKeys & KK <= -(range shrK) \ + (\SK KK. SK \ symKeys \ KK \ -(range shrK) \ (\K \ KK. \ AKcryptSK K SK evs) \ - (Key SK \ analz (Key`KK Un (spies evs))) = + (Key SK \ analz (Key`KK \ (spies evs))) = (SK \ KK | Key SK \ analz (spies evs)))" apply (erule kerbIV.induct) apply (frule_tac [10] Oops_range_spies2) @@ -1213,7 +1213,7 @@ text\First simplification law for analz: no session keys encrypt authentication keys or shared keys.\ lemma analz_insert_freshK1: - "\ evs \ kerbIV; K \ authKeys evs Un range shrK; + "\ evs \ kerbIV; K \ authKeys evs \ range shrK; SesKey \ range shrK \ \ (Key K \ analz (insert (Key SesKey) (spies evs))) = (K = SesKey | Key K \ analz (spies evs))" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/KerberosIV_Gets.thy --- a/src/HOL/Auth/KerberosIV_Gets.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Thu Feb 15 13:04:36 2018 +0100 @@ -22,7 +22,7 @@ definition (* authKeys are those contained in an authTicket *) - authKeys :: "event list => key set" where + authKeys :: "event list \ key set" where "authKeys evs = {authK. \A Peer Ta. Says Kas A (Crypt (shrK A) \Key authK, Agent Peer, Number Ta, (Crypt (shrK Peer) \Agent A, Agent Peer, Key authK, Number Ta\) @@ -30,8 +30,8 @@ definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] => bool" ("Unique _ on _" [0, 50] 50) - where "(Unique ev on evs) = (ev \ set (tl (dropWhile (% z. z \ ev) evs)))" + Unique :: "[event, event list] \ bool" ("Unique _ on _" [0, 50] 50) + where "(Unique ev on evs) = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" consts @@ -65,30 +65,30 @@ abbreviation (*The current time is just the length of the trace!*) - CT :: "event list=>nat" where + CT :: "event list \ nat" where "CT == length" abbreviation - expiredAK :: "[nat, event list] => bool" where + expiredAK :: "[nat, event list] \ bool" where "expiredAK Ta evs == authKlife + Ta < CT evs" abbreviation - expiredSK :: "[nat, event list] => bool" where + expiredSK :: "[nat, event list] \ bool" where "expiredSK Ts evs == servKlife + Ts < CT evs" abbreviation - expiredA :: "[nat, event list] => bool" where + expiredA :: "[nat, event list] \ bool" where "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] => bool" ("valid _ wrt _" [0, 50] 50) where - "valid T1 wrt T2 == T1 <= replylife + T2" + valid :: "[nat, nat] \ bool" ("valid _ wrt _" [0, 50] 50) where + "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) (* Predicate formalising the association between authKeys and servKeys *) -definition AKcryptSK :: "[key, key, event list] => bool" where +definition AKcryptSK :: "[key, key, event list] \ bool" where "AKcryptSK authK servK evs == \A B Ts. Says Tgs A (Crypt authK @@ -164,7 +164,7 @@ \ set evs4; \ expiredAK Ta evs4; \ expiredA T2 evs4; - servKlife + (CT evs4) <= authKlife + Ta + servKlife + (CT evs4) \ authKlife + Ta \ \ Says Tgs A (Crypt authK \Key servK, Agent B, Number (CT evs4), @@ -311,7 +311,7 @@ lemma Oops_range_spies1: "\ Says Kas A (Crypt KeyA \Key authK, Peer, Ta, authTicket\) \ set evs ; - evs \ kerbIV_gets \ \ authK \ range shrK & authK \ symKeys" + evs \ kerbIV_gets \ \ authK \ range shrK \ authK \ symKeys" apply (erule rev_mp) apply (erule kerbIV_gets.induct, auto) done @@ -319,7 +319,7 @@ lemma Oops_range_spies2: "\ Says Tgs A (Crypt authK \Key servK, Agent B, Ts, servTicket\) \ set evs ; - evs \ kerbIV_gets \ \ servK \ range shrK & servK \ symKeys" + evs \ kerbIV_gets \ \ servK \ range shrK \ servK \ symKeys" apply (erule rev_mp) apply (erule kerbIV_gets.induct, auto) done @@ -339,7 +339,7 @@ by auto lemma Spy_see_shrK_D [dest!]: - "\ Key (shrK A) \ parts (spies evs); evs \ kerbIV_gets \ \ A:bad" + "\ Key (shrK A) \ parts (spies evs); evs \ kerbIV_gets \ \ A\bad" by (blast dest: Spy_see_shrK) lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] @@ -374,8 +374,8 @@ "\ Says Kas A (Crypt K \Key authK, Agent Peer, Number Ta, authTicket\) \ set evs; evs \ kerbIV_gets \ \ - K = shrK A & Peer = Tgs & - authK \ range shrK & authK \ authKeys evs & authK \ symKeys & + K = shrK A \ Peer = Tgs \ + authK \ range shrK \ authK \ authKeys evs \ authK \ symKeys \ authTicket = (Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\)" apply (erule rev_mp) apply (erule kerbIV_gets.induct) @@ -424,9 +424,9 @@ "\ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, servTicket\) \ set evs; evs \ kerbIV_gets \ - \ B \ Tgs & - authK \ range shrK & authK \ authKeys evs & authK \ symKeys & - servK \ range shrK & servK \ authKeys evs & servK \ symKeys & + \ B \ Tgs \ + authK \ range shrK \ authK \ authKeys evs \ authK \ symKeys \ + servK \ range shrK \ servK \ authKeys evs \ servK \ symKeys \ servTicket = (Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\)" apply (erule rev_mp) apply (erule kerbIV_gets.induct) @@ -443,7 +443,7 @@ \ parts (spies evs); A \ bad; evs \ kerbIV_gets \ - \ authK \ range shrK & authK \ symKeys & + \ authK \ range shrK \ authK \ symKeys \ authTicket = Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Ta\" apply (erule rev_mp) apply (erule kerbIV_gets.induct) @@ -458,7 +458,7 @@ \ parts (spies evs); Key authK \ analz (spies evs); evs \ kerbIV_gets \ - \ servK \ range shrK & servK \ symKeys & + \ servK \ range shrK \ servK \ symKeys \ (\A. servTicket = Crypt (shrK B) \Agent A, Agent B, Key servK, Ts\)" apply (erule rev_mp) apply (erule rev_mp) @@ -472,7 +472,7 @@ "\ Gets A (Crypt (shrK A) \Key authK, Agent Tgs, Ta, authTicket\) \ set evs; evs \ kerbIV_gets \ - \ authK \ range shrK & authK \ symKeys & + \ authK \ range shrK \ authK \ symKeys \ authTicket = Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Ta\ | authTicket \ analz (spies evs)" @@ -483,7 +483,7 @@ "\ Gets A (Crypt authK \Key servK, Agent B, Ts, servTicket\) \ set evs; authK \ symKeys; evs \ kerbIV_gets \ - \ servK \ range shrK & + \ servK \ range shrK \ (\A. servTicket = Crypt (shrK B) \Agent A, Agent B, Key servK, Ts\) | servTicket \ analz (spies evs)" @@ -593,7 +593,7 @@ \ \Ta. (Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & servKlife + Ts <= authKlife + Ta)" + \ servKlife + Ts \ authKlife + Ta)" apply (erule rev_mp) apply (erule kerbIV_gets.induct) apply (frule_tac [8] Gets_ticket_parts) @@ -619,7 +619,7 @@ \ \authK Ta. Says Kas A (Crypt(shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & servKlife + Ts <= authKlife + Ta" + \ servKlife + Ts \ authKlife + Ta" by (blast dest!: servTicket_authentic_Tgs u_K4_imp_K2) lemma servTicket_authentic: @@ -630,7 +630,7 @@ Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, + \ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\\) \ set evs" by (blast dest: servTicket_authentic_Tgs K4_imp_K2) @@ -643,14 +643,14 @@ (Says Kas A (Crypt (shrK A) \Key authK, Agent Tgs, Number Ta, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\\) \ set evs - & Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, + \ Says Tgs A (Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\\) \ set evs - & servKlife + Ts <= authKlife + Ta)" + \ servKlife + Ts \ authKlife + Ta)" by (blast dest: servTicket_authentic_Tgs u_K4_imp_K2) lemma u_NotexpiredSK_NotexpiredAK: - "\ \ expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \ + "\ \ expiredSK Ts evs; servKlife + Ts \ authKlife + Ta \ \ \ expiredAK Ta evs" by (blast dest: leI le_trans dest: leD) @@ -679,7 +679,7 @@ Crypt K' \Key SesKey, Agent B', T', Ticket'\ \ parts (spies evs); Key SesKey \ analz (spies evs); evs \ kerbIV_gets \ - \ K=K' & B=B' & T=T' & Ticket=Ticket'" + \ K=K' \ B=B' \ T=T' \ Ticket=Ticket'" apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -753,7 +753,7 @@ Crypt (shrK B') \Agent A', Agent B', Key SesKey, T'\ \ parts (spies evs); Key SesKey \ analz (spies evs); evs \ kerbIV_gets \ - \ A=A' & B=B' & T=T'" + \ A=A' \ B=B' \ T=T'" apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -819,7 +819,7 @@ (Crypt Ka \Key authK, Agent Tgs, Ta, X\) \ set evs; Says Kas A' (Crypt Ka' \Key authK, Agent Tgs, Ta', X'\) \ set evs; - evs \ kerbIV_gets \ \ A=A' & Ka=Ka' & Ta=Ta' & X=X'" + evs \ kerbIV_gets \ \ A=A' \ Ka=Ka' \ Ta=Ta' \ X=X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule kerbIV_gets.induct) @@ -835,7 +835,7 @@ (Crypt K \Key servK, Agent B, Ts, X\) \ set evs; Says Tgs A' (Crypt K' \Key servK, Agent B', Ts', X'\) \ set evs; - evs \ kerbIV_gets \ \ A=A' & B=B' & K=K' & Ts=Ts' & X=X'" + evs \ kerbIV_gets \ \ A=A' \ B=B' \ K=K' \ Ts=Ts' \ X=X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule kerbIV_gets.induct) @@ -882,7 +882,7 @@ lemma AKcryptSK_Says [simp]: "AKcryptSK authK servK (Says S A X # evs) = - (Tgs = S & + (Tgs = S \ (\B Ts. X = Crypt authK \Key servK, Agent B, Number Ts, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\ \) @@ -996,9 +996,9 @@ text\We take some pains to express the property as a logical equivalence so that the simplifier can apply it.\ lemma Key_analz_image_Key_lemma: - "P \ (Key K \ analz (Key`KK Un H)) \ (K:KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K \ KK | Key K \ analz H) \ - P \ (Key K \ analz (Key`KK Un H)) = (K:KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K \ KK | Key K \ analz H)" by (blast intro: analz_mono [THEN subsetD]) @@ -1009,7 +1009,7 @@ by (drule Says_imp_spies [THEN analz.Inj, THEN analz_insertI], auto) lemma authKeys_are_not_AKcryptSK: - "\ K \ authKeys evs Un range shrK; evs \ kerbIV_gets \ + "\ K \ authKeys evs \ range shrK; evs \ kerbIV_gets \ \ \SK. \ AKcryptSK SK K evs \ K \ symKeys" apply (simp add: authKeys_def AKcryptSK_def) by (blast dest: Says_Kas_message_form Says_Tgs_message_form) @@ -1039,9 +1039,9 @@ in case of loss of a key to the spy. See ESORICS98.\ lemma Key_analz_image_Key [rule_format (no_asm)]: "evs \ kerbIV_gets \ - (\SK KK. SK \ symKeys & KK <= -(range shrK) \ + (\SK KK. SK \ symKeys \ KK \ -(range shrK) \ (\K \ KK. \ AKcryptSK K SK evs) \ - (Key SK \ analz (Key`KK Un (spies evs))) = + (Key SK \ analz (Key`KK \ (spies evs))) = (SK \ KK | Key SK \ analz (spies evs)))" apply (erule kerbIV_gets.induct) apply (frule_tac [11] Oops_range_spies2) @@ -1084,7 +1084,7 @@ text\First simplification law for analz: no session keys encrypt authentication keys or shared keys.\ lemma analz_insert_freshK1: - "\ evs \ kerbIV_gets; K \ authKeys evs Un range shrK; + "\ evs \ kerbIV_gets; K \ authKeys evs \ range shrK; SesKey \ range shrK \ \ (Key K \ analz (insert (Key SesKey) (spies evs))) = (K = SesKey | Key K \ analz (spies evs))" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/KerberosV.thy --- a/src/HOL/Auth/KerberosV.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/KerberosV.thy Thu Feb 15 13:04:36 2018 +0100 @@ -23,7 +23,7 @@ definition (* authKeys are those contained in an authTicket *) - authKeys :: "event list => key set" where + authKeys :: "event list \ key set" where "authKeys evs = {authK. \A Peer Ta. Says Kas A \Crypt (shrK A) \Key authK, Agent Peer, Ta\, Crypt (shrK Peer) \Agent A, Agent Peer, Key authK, Ta\ @@ -32,11 +32,11 @@ definition (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) - Issues :: "[agent, agent, msg, event list] => bool" + Issues :: "[agent, agent, msg, event list] \ bool" ("_ Issues _ with _ on _") where "A Issues B with X on evs = (\Y. Says A B Y \ set evs \ X \ parts {Y} \ - X \ parts (spies (takeWhile (% z. z \ Says A B Y) (rev evs))))" + X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" consts @@ -70,30 +70,30 @@ abbreviation (*The current time is just the length of the trace!*) - CT :: "event list=>nat" where + CT :: "event list \ nat" where "CT == length" abbreviation - expiredAK :: "[nat, event list] => bool" where + expiredAK :: "[nat, event list] \ bool" where "expiredAK T evs == authKlife + T < CT evs" abbreviation - expiredSK :: "[nat, event list] => bool" where + expiredSK :: "[nat, event list] \ bool" where "expiredSK T evs == servKlife + T < CT evs" abbreviation - expiredA :: "[nat, event list] => bool" where + expiredA :: "[nat, event list] \ bool" where "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] => bool" ("valid _ wrt _") where - "valid T1 wrt T2 == T1 <= replylife + T2" + valid :: "[nat, nat] \ bool" ("valid _ wrt _") where + "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) (* Predicate formalising the association between authKeys and servKeys *) -definition AKcryptSK :: "[key, key, event list] => bool" where +definition AKcryptSK :: "[key, key, event list] \ bool" where "AKcryptSK authK servK evs == \A B tt. Says Tgs A \Crypt authK \Key servK, Agent B, tt\, @@ -142,7 +142,7 @@ \ set evs4; \ expiredAK Ta evs4; \ expiredA T2 evs4; - servKlife + (CT evs4) <= authKlife + Ta + servKlife + (CT evs4) \ authKlife + Ta \ \ Says Tgs A \ Crypt authK \Key servK, Agent B, Number (CT evs4)\, @@ -218,7 +218,7 @@ done lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = - (if A:bad then insert X (spies evs) else spies evs)" + (if A\bad then insert X (spies evs) else spies evs)" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -233,7 +233,7 @@ lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] -lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" +lemma spies_takeWhile: "spies (takeWhile P evs) \ spies evs" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -294,7 +294,7 @@ lemma Oops_range_spies1: "\ Says Kas A \Crypt KeyA \Key authK, Peer, Ta\, authTicket\ \ set evs ; - evs \ kerbV \ \ authK \ range shrK & authK \ symKeys" + evs \ kerbV \ \ authK \ range shrK \ authK \ symKeys" apply (erule rev_mp) apply (erule kerbV.induct, auto) done @@ -322,7 +322,7 @@ by auto lemma Spy_see_shrK_D [dest!]: - "\ Key (shrK A) \ parts (spies evs); evs \ kerbV \ \ A:bad" + "\ Key (shrK A) \ parts (spies evs); evs \ kerbV \ \ A\bad" by (blast dest: Spy_see_shrK) lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] @@ -527,7 +527,7 @@ \ \Ta. Says Kas A \Crypt (shrK A) \Key authK, Agent Tgs, Number Ta\, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\ \ \ set evs - \ servKlife + Ts <= authKlife + Ta" + \ servKlife + Ts \ authKlife + Ta" apply (erule rev_mp) apply (erule kerbV.induct) apply (frule_tac [7] Says_ticket_parts) @@ -555,7 +555,7 @@ \Crypt (shrK A) \Key authK, Agent Tgs, Number Ta\, Crypt (shrK Tgs) \Agent A, Agent Tgs, Key authK, Number Ta\ \ \ set evs \ - servKlife + Ts <= authKlife + Ta" + servKlife + Ts \ authKlife + Ta" by (metis servTicket_authentic_Tgs u_K4_imp_K2) lemma servTicket_authentic: @@ -580,11 +580,11 @@ \ Says Tgs A \Crypt authK \Key servK, Agent B, Number Ts\, Crypt (shrK B) \Agent A, Agent B, Key servK, Number Ts\\ \ set evs - \ servKlife + Ts <= authKlife + Ta" + \ servKlife + Ts \ authKlife + Ta" by (metis servTicket_authentic_Tgs u_K4_imp_K2) lemma u_NotexpiredSK_NotexpiredAK: - "\ \ expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \ + "\ \ expiredSK Ts evs; servKlife + Ts \ authKlife + Ta \ \ \ expiredAK Ta evs" by (metis order_le_less_trans) @@ -653,7 +653,7 @@ Crypt (shrK B') \Agent A', Agent B', Key SesKey, T'\ \ parts (spies evs); Key SesKey \ analz (spies evs); evs \ kerbV \ - \ A=A' & B=B' & T=T'" + \ A=A' \ B=B' \ T=T'" apply (erule rev_mp) apply (erule rev_mp) apply (erule rev_mp) @@ -868,9 +868,9 @@ text\We take some pains to express the property as a logical equivalence so that the simplifier can apply it.\ lemma Key_analz_image_Key_lemma: - "P \ (Key K \ analz (Key`KK Un H)) \ (K:KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K\KK \ Key K \ analz H) \ - P \ (Key K \ analz (Key`KK Un H)) = (K:KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K\KK \ Key K \ analz H)" by (blast intro: analz_mono [THEN subsetD]) @@ -882,7 +882,7 @@ done lemma authKeys_are_not_AKcryptSK: - "\ K \ authKeys evs Un range shrK; evs \ kerbV \ + "\ K \ authKeys evs \ range shrK; evs \ kerbV \ \ \SK. \ AKcryptSK SK K evs \ K \ symKeys" apply (simp add: authKeys_def AKcryptSK_def) apply (blast dest: Says_Kas_message_form Says_Tgs_message_form) @@ -914,9 +914,9 @@ in case of loss of a key to the spy. See ESORICS98.\ lemma Key_analz_image_Key [rule_format (no_asm)]: "evs \ kerbV \ - (\SK KK. SK \ symKeys & KK <= -(range shrK) \ + (\SK KK. SK \ symKeys \ KK \ -(range shrK) \ (\K \ KK. \ AKcryptSK K SK evs) \ - (Key SK \ analz (Key`KK Un (spies evs))) = + (Key SK \ analz (Key`KK \ (spies evs))) = (SK \ KK | Key SK \ analz (spies evs)))" apply (erule kerbV.induct) apply (frule_tac [10] Oops_range_spies2) @@ -951,7 +951,7 @@ text\First simplification law for analz: no session keys encrypt authentication keys or shared keys.\ lemma analz_insert_freshK1: - "\ evs \ kerbV; K \ authKeys evs Un range shrK; + "\ evs \ kerbV; K \ authKeys evs \ range shrK; SesKey \ range shrK \ \ (Key K \ analz (insert (Key SesKey) (spies evs))) = (K = SesKey | Key K \ analz (spies evs))" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Kerberos_BAN.thy --- a/src/HOL/Auth/Kerberos_BAN.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Kerberos_BAN.thy Thu Feb 15 13:04:36 2018 +0100 @@ -38,36 +38,36 @@ by blast abbreviation - CT :: "event list=>nat" where + CT :: "event list \ nat" where "CT == length " abbreviation - expiredK :: "[nat, event list] => bool" where + expiredK :: "[nat, event list] \ bool" where "expiredK T evs == sesKlife + T < CT evs" abbreviation - expiredA :: "[nat, event list] => bool" where + expiredA :: "[nat, event list] \ bool" where "expiredA T evs == authlife + T < CT evs" definition (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) - Issues :: "[agent, agent, msg, event list] => bool" + Issues :: "[agent, agent, msg, event list] \ bool" ("_ Issues _ with _ on _") where "A Issues B with X on evs = - (\Y. Says A B Y \ set evs & X \ parts {Y} & - X \ parts (spies (takeWhile (% z. z \ Says A B Y) (rev evs))))" + (\Y. Says A B Y \ set evs \ X \ parts {Y} \ + X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] => event list" ("before _ on _") - where "before ev on evs = takeWhile (% z. z ~= ev) (rev evs)" + before :: "[event, event list] \ event list" ("before _ on _") + where "before ev on evs = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] => bool" ("Unique _ on _") - where "Unique ev on evs = (ev \ set (tl (dropWhile (% z. z \ ev) evs)))" + Unique :: "[event, event list] \ bool" ("Unique _ on _") + where "Unique ev on evs = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" inductive_set bankerberos :: "event list set" @@ -104,7 +104,7 @@ | BK4: "\ evs4 \ bankerberos; Says A' B \(Crypt (shrK B) \Number Tk, Agent A, Key K\), - (Crypt K \Agent A, Number Ta\) \: set evs4; + (Crypt K \Agent A, Number Ta\) \ \ set evs4; \ expiredK Tk evs4; \ expiredA Ta evs4 \ \ Says B A (Crypt K (Number Ta)) # evs4 \ bankerberos" @@ -150,7 +150,7 @@ done lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = - (if A:bad then insert X (spies evs) else spies evs)" + (if A\bad then insert X (spies evs) else spies evs)" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -165,7 +165,7 @@ lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] -lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" +lemma spies_takeWhile: "spies (takeWhile P evs) \ spies evs" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -211,7 +211,7 @@ done lemma used_takeWhile_used [rule_format]: - "x : used (takeWhile P X) --> x : used X" + "x \ used (takeWhile P X) \ x \ used X" apply (induct_tac "X") apply simp apply (rename_tac a b) @@ -260,7 +260,7 @@ lemma Spy_see_shrK_D [dest!]: "\ Key (shrK A) \ parts (spies evs); - evs \ bankerberos \ \ A:bad" + evs \ bankerberos \ \ A\bad" apply (blast dest: Spy_see_shrK) done @@ -287,11 +287,11 @@ lemma Says_Server_message_form: "\ Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) \ set evs; evs \ bankerberos \ - \ K' = shrK A & K \ range shrK & - Ticket = (Crypt (shrK B) \Number Tk, Agent A, Key K\) & + \ K' = shrK A \ K \ range shrK \ + Ticket = (Crypt (shrK B) \Number Tk, Agent A, Key K\) \ Key K \ used(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) - on evs) & + on evs) \ Tk = CT(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) on evs)" @@ -343,7 +343,7 @@ "\ Says S A (Crypt (shrK A) \Number Tk, Agent B, Key K, X\) \ set evs; evs \ bankerberos \ - \ (K \ range shrK & X = (Crypt (shrK B) \Number Tk, Agent A, Key K\)) + \ (K \ range shrK \ X = (Crypt (shrK B) \Number Tk, Agent A, Key K\)) | X \ analz (spies evs)" apply (case_tac "A \ bad") apply (force dest!: Says_imp_spies [THEN analz.Inj]) @@ -367,7 +367,7 @@ lemma analz_image_freshK [rule_format (no_asm)]: "evs \ bankerberos \ \K KK. KK \ - (range shrK) \ - (Key K \ analz (Key`KK Un (spies evs))) = + (Key K \ analz (Key`KK \ (spies evs))) = (K \ KK | Key K \ analz (spies evs))" apply (erule bankerberos.induct) apply (drule_tac [7] Says_Server_message_form) @@ -388,7 +388,7 @@ (Crypt (shrK A) \Number Tk, Agent B, Key K, X\) \ set evs; Says Server A' (Crypt (shrK A') \Number Tk', Agent B', Key K, X'\) \ set evs; - evs \ bankerberos \ \ A=A' & Tk=Tk' & B=B' & X = X'" + evs \ bankerberos \ \ A=A' \ Tk=Tk' \ B=B' \ X = X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule bankerberos.induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Kerberos_BAN_Gets.thy --- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Thu Feb 15 13:04:36 2018 +0100 @@ -40,27 +40,27 @@ abbreviation - CT :: "event list=>nat" where + CT :: "event list \ nat" where "CT == length" abbreviation - expiredK :: "[nat, event list] => bool" where + expiredK :: "[nat, event list] \ bool" where "expiredK T evs == sesKlife + T < CT evs" abbreviation - expiredA :: "[nat, event list] => bool" where + expiredA :: "[nat, event list] \ bool" where "expiredA T evs == authlife + T < CT evs" definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] => event list" ("before _ on _") - where "before ev on evs = takeWhile (% z. z ~= ev) (rev evs)" + before :: "[event, event list] \ event list" ("before _ on _") + where "before ev on evs = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] => bool" ("Unique _ on _") - where "Unique ev on evs = (ev \ set (tl (dropWhile (% z. z \ ev) evs)))" + Unique :: "[event, event list] \ bool" ("Unique _ on _") + where "Unique ev on evs = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" inductive_set bankerb_gets :: "event list set" @@ -99,7 +99,7 @@ | BK4: "\ evs4 \ bankerb_gets; Gets B \(Crypt (shrK B) \Number Tk, Agent A, Key K\), - (Crypt K \Agent A, Number Ta\) \: set evs4; + (Crypt K \Agent A, Number Ta\) \ \ set evs4; \ expiredK Tk evs4; \ expiredA Ta evs4 \ \ Says B A (Crypt K (Number Ta)) # evs4 \ bankerb_gets" @@ -200,7 +200,7 @@ done lemma used_takeWhile_used [rule_format]: - "x : used (takeWhile P X) --> x : used X" + "x \ used (takeWhile P X) \ x \ used X" apply (induct_tac "X") apply simp apply (rename_tac a b) @@ -249,7 +249,7 @@ lemma Spy_see_shrK_D [dest!]: "\ Key (shrK A) \ parts (knows Spy evs); - evs \ bankerb_gets \ \ A:bad" + evs \ bankerb_gets \ \ A\bad" by (blast dest: Spy_see_shrK) lemmas Spy_analz_shrK_D = analz_subset_parts [THEN subsetD, THEN Spy_see_shrK_D, dest!] @@ -275,11 +275,11 @@ lemma Says_Server_message_form: "\ Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) \ set evs; evs \ bankerb_gets \ - \ K' = shrK A & K \ range shrK & - Ticket = (Crypt (shrK B) \Number Tk, Agent A, Key K\) & + \ K' = shrK A \ K \ range shrK \ + Ticket = (Crypt (shrK B) \Number Tk, Agent A, Key K\) \ Key K \ used(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) - on evs) & + on evs) \ Tk = CT(before Says Server A (Crypt K' \Number Tk, Agent B, Key K, Ticket\) on evs)" @@ -337,7 +337,7 @@ "\ Gets A (Crypt (shrK A) \Number Tk, Agent B, Key K, X\) \ set evs; evs \ bankerb_gets \ - \ (K \ range shrK & X = (Crypt (shrK B) \Number Tk, Agent A, Key K\)) + \ (K \ range shrK \ X = (Crypt (shrK B) \Number Tk, Agent A, Key K\)) | X \ analz (knows Spy evs)" apply (case_tac "A \ bad") apply (force dest!: Gets_imp_knows_Spy [THEN analz.Inj]) @@ -398,7 +398,7 @@ lemma analz_image_freshK [rule_format (no_asm)]: "evs \ bankerb_gets \ \K KK. KK \ - (range shrK) \ - (Key K \ analz (Key`KK Un (knows Spy evs))) = + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule bankerb_gets.induct) apply (drule_tac [8] Says_Server_message_form) @@ -419,7 +419,7 @@ (Crypt (shrK A) \Number Tk, Agent B, Key K, X\) \ set evs; Says Server A' (Crypt (shrK A') \Number Tk', Agent B', Key K, X'\) \ set evs; - evs \ bankerb_gets \ \ A=A' & Tk=Tk' & B=B' & X = X'" + evs \ bankerb_gets \ \ A=A' \ Tk=Tk' \ B=B' \ X = X'" apply (erule rev_mp) apply (erule rev_mp) apply (erule bankerb_gets.induct) @@ -434,7 +434,7 @@ (Crypt (shrK A) \Number Tk, Agent B, Key K, X\) \ set evs; Gets A (Crypt (shrK A) \Number Tk', Agent B', Key K, X'\) \ set evs; - A \ bad; evs \ bankerb_gets \ \ Tk=Tk' & B=B' & X = X'" + A \ bad; evs \ bankerb_gets \ \ Tk=Tk' \ B=B' \ X = X'" apply (blast dest!: Kab_authentic unique_session_keys) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Message.thy Thu Feb 15 13:04:36 2018 +0100 @@ -21,11 +21,11 @@ consts all_symmetric :: bool \ \true if all keys are symmetric\ - invKey :: "key=>key" \ \inverse of a symmetric key\ + invKey :: "key\key" \ \inverse of a symmetric key\ specification (invKey) invKey [simp]: "invKey (invKey K) = K" - invKey_symmetric: "all_symmetric --> invKey = id" + invKey_symmetric: "all_symmetric \ invKey = id" by (rule exI [of _ id], auto) @@ -56,11 +56,11 @@ "\x, y\" \ "CONST MPair x y" -definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where +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\" -definition keysFor :: "msg set => key set" where +definition keysFor :: "msg set \ key set" where \ \Keys useful to decrypt elements of a message set\ "keysFor H == invKey ` {K. \X. Crypt K X \ H}" @@ -68,7 +68,7 @@ subsubsection\Inductive Definition of All Parts" of a Message\ inductive_set - parts :: "msg set => msg set" + parts :: "msg set \ msg set" for H :: "msg set" where Inj [intro]: "X \ H ==> X \ parts H" @@ -86,7 +86,7 @@ text\Equations hold because constructors are injective.\ -lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" +lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x\A)" by auto lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x\A)" @@ -310,7 +310,7 @@ by auto text\In any message, there is an upper bound N on its greatest nonce.\ -lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" +lemma msg_Nonce_supply: "\N. \n. N\n \ Nonce n \ parts {msg}" proof (induct msg) case (Nonce n) show ?case @@ -328,14 +328,14 @@ be taken apart; messages decrypted with known keys.\ inductive_set - analz :: "msg set => msg set" + analz :: "msg set \ msg set" 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" | Decrypt [dest]: - "[|Crypt K X \ analz H; Key(invKey K): analz H|] ==> X \ analz H" + "\Crypt K X \ analz H; Key(invKey K) \ analz H\ \ X \ analz H" text\Monotonicity; Lemma 1 of Lowe's paper\ @@ -513,7 +513,7 @@ by (erule analz_trans, blast) (*Cut can be proved easily by induction on - "Y: analz (insert X H) ==> X: analz H --> Y: analz H" + "Y: analz (insert X H) ==> X: analz H \ Y: analz H" *) text\This rewrite rule helps in the simplification of messages that involve @@ -670,8 +670,8 @@ parts_synth synth_mono synth_subset_iff) lemma Fake_parts_insert_in_Un: - "[|Z \ parts (insert X H); X: synth (analz H)|] - ==> Z \ synth (analz H) \ parts H" + "\Z \ parts (insert X H); X \ synth (analz H)\ + \ Z \ synth (analz H) \ parts H" by (metis Fake_parts_insert set_mp) text\@{term H} is sometimes @{term"Key ` KK \ spies evs"}, so can't put @@ -685,7 +685,7 @@ done lemma analz_conj_parts [simp]: - "(X \ analz H & X \ parts H) = (X \ analz H)" + "(X \ analz H \ X \ parts H) = (X \ analz H)" by (blast intro: analz_subset_parts [THEN subsetD]) lemma analz_disj_parts [simp]: @@ -696,7 +696,7 @@ redundant cases\ lemma MPair_synth_analz [iff]: "(\X,Y\ \ synth (analz H)) = - (X \ synth (analz H) & Y \ synth (analz H))" + (X \ synth (analz H) \ Y \ synth (analz H))" by blast lemma Crypt_synth_analz: @@ -715,22 +715,22 @@ subsubsection\Freeness\ -lemma Agent_neq_HPair: "Agent A ~= Hash[X] Y" +lemma Agent_neq_HPair: "Agent A \ Hash[X] Y" unfolding HPair_def by simp -lemma Nonce_neq_HPair: "Nonce N ~= Hash[X] Y" +lemma Nonce_neq_HPair: "Nonce N \ Hash[X] Y" unfolding HPair_def by simp -lemma Number_neq_HPair: "Number N ~= Hash[X] Y" +lemma Number_neq_HPair: "Number N \ Hash[X] Y" unfolding HPair_def by simp -lemma Key_neq_HPair: "Key K ~= Hash[X] Y" +lemma Key_neq_HPair: "Key K \ Hash[X] Y" unfolding HPair_def by simp -lemma Hash_neq_HPair: "Hash Z ~= Hash[X] Y" +lemma Hash_neq_HPair: "Hash Z \ Hash[X] Y" unfolding HPair_def by simp -lemma Crypt_neq_HPair: "Crypt K X' ~= Hash[X] Y" +lemma Crypt_neq_HPair: "Crypt K X' \ Hash[X] Y" unfolding HPair_def by simp lemmas HPair_neqs = Agent_neq_HPair Nonce_neq_HPair Number_neq_HPair @@ -739,15 +739,15 @@ declare HPair_neqs [iff] declare HPair_neqs [symmetric, iff] -lemma HPair_eq [iff]: "(Hash[X'] Y' = Hash[X] Y) = (X' = X & Y'=Y)" +lemma HPair_eq [iff]: "(Hash[X'] Y' = Hash[X] Y) = (X' = X \ Y'=Y)" by (simp add: HPair_def) lemma MPair_eq_HPair [iff]: - "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ & Y'=Y)" + "(\X',Y'\ = Hash[X] Y) = (X' = Hash\X,Y\ \ Y'=Y)" by (simp add: HPair_def) lemma HPair_eq_MPair [iff]: - "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ & Y'=Y)" + "(Hash[X] Y = \X',Y'\) = (X' = Hash\X,Y\ \ Y'=Y)" by (auto simp add: HPair_def) @@ -769,7 +769,7 @@ lemma HPair_synth_analz [simp]: "X \ synth (analz H) ==> (Hash[X] Y \ synth (analz H)) = - (Hash \X, Y\ \ analz H & Y \ synth (analz H))" + (Hash \X, Y\ \ analz H \ Y \ synth (analz H))" by (auto simp add: HPair_def) @@ -897,12 +897,12 @@ text\Two generalizations of \analz_insert_eq\\ lemma gen_analz_insert_eq [rule_format]: - "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G" + "X \ analz H \ \G. H \ G \ analz (insert X G) = analz G" by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) lemma synth_analz_insert_eq [rule_format]: "X \ synth (analz H) - ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)" + \ \G. H \ G \ (Key K \ analz (insert X G)) = (Key K \ analz G)" apply (erule synth.induct) apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/NS_Shared.thy --- a/src/HOL/Auth/NS_Shared.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/NS_Shared.thy Thu Feb 15 13:04:36 2018 +0100 @@ -16,11 +16,11 @@ definition (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) - Issues :: "[agent, agent, msg, event list] => bool" + Issues :: "[agent, agent, msg, event list] \ bool" ("_ Issues _ with _ on _") where "A Issues B with X on evs = - (\Y. Says A B Y \ set evs & X \ parts {Y} & - X \ parts (spies (takeWhile (% z. z \ Says A B Y) (rev evs))))" + (\Y. Says A B Y \ set evs \ X \ parts {Y} \ + X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" inductive_set ns_shared :: "event list set" @@ -395,7 +395,7 @@ done lemma spies_Notes_rev: "spies (evs @ [Notes A X]) = - (if A:bad then insert X (spies evs) else spies evs)" + (if A\bad then insert X (spies evs) else spies evs)" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) @@ -410,7 +410,7 @@ lemmas parts_spies_evs_revD2 = spies_evs_rev [THEN equalityD2, THEN parts_mono] -lemma spies_takeWhile: "spies (takeWhile P evs) <= spies evs" +lemma spies_takeWhile: "spies (takeWhile P evs) \ spies evs" apply (induct_tac "evs") apply (rename_tac [2] a b) apply (induct_tac [2] "a", auto) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/OtwayRees.thy --- a/src/HOL/Auth/OtwayRees.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/OtwayRees.thy Thu Feb 15 13:04:36 2018 +0100 @@ -37,7 +37,7 @@ (*Bob's response to Alice's message. Note that NB is encrypted.*) | OR2: "\evs2 \ otway; Nonce NB \ used evs2; - Gets B \Nonce NA, Agent A, Agent B, X\ : set evs2\ + Gets B \Nonce NA, Agent A, Agent B, X\ \ set evs2\ \ Says B Server \Nonce NA, Agent A, Agent B, X, Crypt (shrK B) @@ -52,7 +52,7 @@ \Nonce NA, Agent A, Agent B, Crypt (shrK A) \Nonce NA, Agent A, Agent B\, Crypt (shrK B) \Nonce NA, Nonce NB, Agent A, Agent B\\ - : set evs3\ + \ set evs3\ \ Says Server B \Nonce NA, Crypt (shrK A) \Nonce NA, Key KAB\, @@ -66,16 +66,16 @@ Says B Server \Nonce NA, Agent A, Agent B, X', Crypt (shrK B) \Nonce NA, Nonce NB, Agent A, Agent B\\ - : set evs4; + \ set evs4; Gets B \Nonce NA, X, Crypt (shrK B) \Nonce NB, Key K\\ - : set evs4\ + \ set evs4\ \ Says B A \Nonce NA, X\ # evs4 \ otway" (*This message models possible leaks of session keys. The nonces identify the protocol run.*) | Oops: "\evso \ otway; Says Server B \Nonce NA, X, Crypt (shrK B) \Nonce NB, Key K\\ - : set evso\ + \ set evso\ \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ otway" @@ -153,7 +153,7 @@ lemma Says_Server_message_form: "\Says Server B \NA, X, Crypt (shrK B) \NB, Key K\\ \ set evs; evs \ otway\ - \ K \ range shrK & (\i. NA = Nonce i) & (\j. NB = Nonce j)" + \ K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" by (erule rev_mp, erule otway.induct, simp_all) @@ -172,8 +172,8 @@ text\The equality makes the induction hypothesis easier to apply\ lemma analz_image_freshK [rule_format]: "evs \ otway \ - \K KK. KK <= -(range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ -(range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule otway.induct) apply (frule_tac [8] Says_Server_message_form) @@ -192,7 +192,7 @@ lemma unique_session_keys: "\Says Server B \NA, X, Crypt (shrK B) \NB, K\\ \ set evs; Says Server B' \NA',X',Crypt (shrK B') \NB',K\\ \ set evs; - evs \ otway\ \ X=X' & B=B' & NA=NA' & NB=NB'" + evs \ otway\ \ X=X' \ B=B' \ NA=NA' \ NB=NB'" apply (erule rev_mp) apply (erule rev_mp) apply (erule otway.induct, simp_all) @@ -205,7 +205,7 @@ text\Only OR1 can have caused such a part of a message to appear.\ lemma Crypt_imp_OR1 [rule_format]: "\A \ bad; evs \ otway\ - \ Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) --> + \ Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) \ Says A B \NA, Agent A, Agent B, Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs" @@ -251,9 +251,9 @@ lemma NA_Crypt_imp_Server_msg [rule_format]: "\A \ bad; evs \ otway\ \ Says A B \NA, Agent A, Agent B, - Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs --> + Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs \ Crypt (shrK A) \NA, Key K\ \ parts (knows Spy evs) - --> (\NB. Says Server B + \ (\NB. Says Server B \NA, Crypt (shrK A) \NA, Key K\, Crypt (shrK B) \NB, Key K\\ \ set evs)" @@ -292,8 +292,8 @@ "\A \ bad; B \ bad; evs \ otway\ \ Says Server B \NA, Crypt (shrK A) \NA, Key K\, - Crypt (shrK B) \NB, Key K\\ \ set evs --> - Notes Spy \NA, NB, Key K\ \ set evs --> + Crypt (shrK B) \NB, Key K\\ \ set evs \ + Notes Spy \NA, NB, Key K\ \ set evs \ Key K \ analz (knows Spy evs)" apply (erule otway.induct, force, simp_all) subgoal \ \Fake\ @@ -368,7 +368,7 @@ "\Crypt (shrK B) \NA, NB, Agent A, Agent B\ \ parts(knows Spy evs); Crypt (shrK B) \NC, NB, Agent C, Agent B\ \ parts(knows Spy evs); evs \ otway; B \ bad\ - \ NC = NA & C = A" + \ NC = NA \ C = A" apply (erule rev_mp, erule rev_mp) apply (erule otway.induct, force, drule_tac [4] OR2_parts_knows_Spy, simp_all) @@ -380,11 +380,11 @@ lemma NB_Crypt_imp_Server_msg [rule_format]: "\B \ bad; evs \ otway\ \ Crypt (shrK B) \NB, Key K\ \ parts (knows Spy evs) - --> (\X'. Says B Server + \ (\X'. Says B Server \NA, Agent A, Agent B, X', Crypt (shrK B) \NA, NB, Agent A, Agent B\\ \ set evs - --> Says Server B + \ Says Server B \NA, Crypt (shrK A) \NA, Key K\, Crypt (shrK B) \NB, Key K\\ \ set evs)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/OtwayReesBella.thy --- a/src/HOL/Auth/OtwayReesBella.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/OtwayReesBella.thy Thu Feb 15 13:04:36 2018 +0100 @@ -166,7 +166,7 @@ lemma Says_Server_message_form: "\Says Server B \Nonce M, Crypt (shrK B) \X, Nonce Nb, Key K\\ \ set evs; evs \ orb\ - \ K \ range shrK & (\ A Na. X=(Crypt (shrK A) \Nonce Na, Key K\))" + \ K \ range shrK \ (\ A Na. X=(Crypt (shrK A) \Nonce Na, Key K\))" by (erule rev_mp, erule orb.induct, simp_all) lemma Says_Server_imp_Gets: @@ -212,19 +212,19 @@ lemma Gets_Server_message_form: "\Gets B \Nonce M, Crypt (shrK B) \X, Nonce Nb, Key K\\ \ set evs; evs \ orb\ - \ (K \ range shrK & (\ A Na. X = (Crypt (shrK A) \Nonce Na, Key K\))) + \ (K \ range shrK \ (\ A Na. X = (Crypt (shrK A) \Nonce Na, Key K\))) | X \ analz (knows Spy evs)" by (metis B_trusts_OR3 Crypt_Spy_analz_bad Gets_imp_Says MPair_analz MPair_parts Says_Server_message_form Says_imp_analz_Spy Says_imp_parts_knows_Spy) lemma unique_Na: "\Says A B \Nonce M, Agent A, Agent B, Crypt (shrK A) \Nonce Na, Nonce M, Agent A, Agent B\\ \ set evs; Says A B' \Nonce M', Agent A, Agent B', Crypt (shrK A) \Nonce Na, Nonce M', Agent A, Agent B'\\ \ set evs; - A \ bad; evs \ orb\ \ B=B' & M=M'" + A \ bad; evs \ orb\ \ B=B' \ M=M'" by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+) lemma unique_Nb: "\Says B Server \Nonce M, Agent A, Agent B, X, Crypt (shrK B) \Nonce Nb, Nonce M, Nonce M, Agent A, Agent B\\ \ set evs; Says B Server \Nonce M', Agent A', Agent B, X', Crypt (shrK B) \Nonce Nb,Nonce M', Nonce M', Agent A', Agent B\\ \ set evs; - B \ bad; evs \ orb\ \ M=M' & A=A' & X=X'" + B \ bad; evs \ orb\ \ M=M' \ A=A' \ X=X'" by (erule rev_mp, erule rev_mp, erule orb.induct, simp_all, blast+) lemma analz_image_freshCryptK_lemma: @@ -342,7 +342,7 @@ lemma Gets_Server_message_form': "\Gets B \Nonce M, Crypt (shrK B) \X, Nonce Nb, Key K\\ \ set evs; B \ bad; evs \ orb\ - \ K \ range shrK & (\ A Na. X = (Crypt (shrK A) \Nonce Na, Key K\))" + \ K \ range shrK \ (\ A Na. X = (Crypt (shrK A) \Nonce Na, Key K\))" by (blast dest!: B_trusts_OR3 Says_Server_message_form) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/OtwayRees_AN.thy --- a/src/HOL/Auth/OtwayRees_AN.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/OtwayRees_AN.thy Thu Feb 15 13:04:36 2018 +0100 @@ -135,7 +135,7 @@ Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs; evs \ otway |] - ==> K \ range shrK & (\i. NA = Nonce i) & (\j. NB = Nonce j)" + ==> K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" apply (erule rev_mp) apply (erule otway.induct, auto) done @@ -157,8 +157,8 @@ text\The equality makes the induction hypothesis easier to apply\ lemma analz_image_freshK [rule_format]: "evs \ otway ==> - \K KK. KK <= -(range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ -(range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule otway.induct) apply (frule_tac [8] Says_Server_message_form) @@ -183,7 +183,7 @@ Crypt (shrK B') \NB', Agent A', Agent B', K\\ \ set evs; evs \ otway |] - ==> A=A' & B=B' & NA=NA' & NB=NB'" + ==> A=A' \ B=B' \ NA=NA' \ NB=NB'" apply (erule rev_mp, erule rev_mp, erule otway.induct, simp_all) apply blast+ \ \OR3 and OR4\ done @@ -195,7 +195,7 @@ lemma NA_Crypt_imp_Server_msg [rule_format]: "[| A \ bad; A \ B; evs \ otway |] ==> Crypt (shrK A) \NA, Agent A, Agent B, Key K\ \ parts (knows Spy evs) - --> (\NB. Says Server B + \ (\NB. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs)" @@ -225,8 +225,8 @@ ==> Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ - \ set evs --> - Notes Spy \NA, NB, Key K\ \ set evs --> + \ set evs \ + Notes Spy \NA, NB, Key K\ \ set evs \ Key K \ analz (knows Spy evs)" apply (erule otway.induct, force) apply (frule_tac [7] Says_Server_message_form) @@ -265,7 +265,7 @@ lemma NB_Crypt_imp_Server_msg [rule_format]: "[| B \ bad; A \ B; evs \ otway |] ==> Crypt (shrK B) \NB, Agent A, Agent B, Key K\ \ parts (knows Spy evs) - --> (\NA. Says Server B + \ (\NA. Says Server B \Crypt (shrK A) \NA, Agent A, Agent B, Key K\, Crypt (shrK B) \NB, Agent A, Agent B, Key K\\ \ set evs)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/OtwayRees_Bad.thy --- a/src/HOL/Auth/OtwayRees_Bad.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/OtwayRees_Bad.thy Thu Feb 15 13:04:36 2018 +0100 @@ -157,7 +157,7 @@ lemma Says_Server_message_form: "[| Says Server B \NA, X, Crypt (shrK B) \NB, Key K\\ \ set evs; evs \ otway |] - ==> K \ range shrK & (\i. NA = Nonce i) & (\j. NB = Nonce j)" + ==> K \ range shrK \ (\i. NA = Nonce i) \ (\j. NB = Nonce j)" apply (erule rev_mp) apply (erule otway.induct, simp_all) done @@ -178,8 +178,8 @@ text\The equality makes the induction hypothesis easier to apply\ lemma analz_image_freshK [rule_format]: "evs \ otway ==> - \K KK. KK <= -(range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ -(range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule otway.induct) apply (frule_tac [8] Says_Server_message_form) @@ -198,7 +198,7 @@ lemma unique_session_keys: "[| Says Server B \NA, X, Crypt (shrK B) \NB, K\\ \ set evs; Says Server B' \NA',X',Crypt (shrK B') \NB',K\\ \ set evs; - evs \ otway |] ==> X=X' & B=B' & NA=NA' & NB=NB'" + evs \ otway |] ==> X=X' \ B=B' \ NA=NA' \ NB=NB'" apply (erule rev_mp) apply (erule rev_mp) apply (erule otway.induct, simp_all) @@ -213,8 +213,8 @@ "[| A \ bad; B \ bad; evs \ otway |] ==> Says Server B \NA, Crypt (shrK A) \NA, Key K\, - Crypt (shrK B) \NB, Key K\\ \ set evs --> - Notes Spy \NA, NB, Key K\ \ set evs --> + Crypt (shrK B) \NB, Key K\\ \ set evs \ + Notes Spy \NA, NB, Key K\ \ set evs \ Key K \ analz (knows Spy evs)" apply (erule otway.induct, force) apply (frule_tac [7] Says_Server_message_form) @@ -243,7 +243,7 @@ up. Original Otway-Rees doesn't need it.\ lemma Crypt_imp_OR1 [rule_format]: "[| A \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) --> + ==> Crypt (shrK A) \NA, Agent A, Agent B\ \ parts (knows Spy evs) \ Says A B \NA, Agent A, Agent B, Crypt (shrK A) \NA, Agent A, Agent B\\ \ set evs" by (erule otway.induct, force, @@ -256,10 +256,10 @@ text\Only it is FALSE. Somebody could make a fake message to Server substituting some other nonce NA' for NB.\ lemma "[| A \ bad; A \ B; evs \ otway |] - ==> Crypt (shrK A) \NA, Key K\ \ parts (knows Spy evs) --> + ==> Crypt (shrK A) \NA, Key K\ \ parts (knows Spy evs) \ Says A B \NA, Agent A, Agent B, Crypt (shrK A) \NA, Agent A, Agent B\\ - \ set evs --> + \ set evs \ (\B NB. Says Server B \NA, Crypt (shrK A) \NA, Key K\, diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Public.thy Thu Feb 15 13:04:36 2018 +0100 @@ -19,27 +19,27 @@ datatype keymode = Signature | Encryption consts - publicKey :: "[keymode,agent] => key" + publicKey :: "[keymode,agent] \ key" abbreviation - pubEK :: "agent => key" where + pubEK :: "agent \ key" where "pubEK == publicKey Encryption" abbreviation - pubSK :: "agent => key" where + pubSK :: "agent \ key" where "pubSK == publicKey Signature" abbreviation - privateKey :: "[keymode, agent] => key" where + privateKey :: "[keymode, agent] \ key" where "privateKey b A == invKey (publicKey b A)" abbreviation (*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*) - priEK :: "agent => key" where + priEK :: "agent \ key" where "priEK A == privateKey Encryption A" abbreviation - priSK :: "agent => key" where + priSK :: "agent \ key" where "priSK A == privateKey Signature A" @@ -47,11 +47,11 @@ simple situation where the signature and encryption keys are the same.\ abbreviation - pubK :: "agent => key" where + pubK :: "agent \ key" where "pubK A == pubEK A" abbreviation - priK :: "agent => key" where + priK :: "agent \ key" where "priK A == invKey (pubEK A)" @@ -59,9 +59,9 @@ @{term "True\False"}, no agent has identical signing and encryption keys\ specification (publicKey) injective_publicKey: - "publicKey b A = publicKey c A' ==> b=c & A=A'" + "publicKey b A = publicKey c A' ==> b=c \ A=A'" apply (rule exI [of _ - "%b A. 2 * case_agent 0 (\n. n + 2) 1 A + case_keymode 0 1 b"]) + "\b A. 2 * case_agent 0 (\n. n + 2) 1 A + case_keymode 0 1 b"]) apply (auto simp add: inj_on_def split: agent.split keymode.split) apply presburger apply presburger @@ -79,7 +79,7 @@ subsection\Basic properties of @{term pubK} and @{term priK}\ -lemma publicKey_inject [iff]: "(publicKey b A = publicKey c A') = (b=c & A=A')" +lemma publicKey_inject [iff]: "(publicKey b A = publicKey c A') = (b=c \ A=A')" by (blast dest!: injective_publicKey) lemma not_symKeys_pubK [iff]: "publicKey b A \ symKeys" @@ -111,14 +111,14 @@ (*holds because invKey is injective*) lemma publicKey_image_eq [simp]: - "(publicKey b x \ publicKey c ` AA) = (b=c & x \ AA)" + "(publicKey b x \ publicKey c ` AA) = (b=c \ x \ AA)" by auto lemma privateKey_notin_image_publicKey [simp]: "privateKey b x \ publicKey c ` AA" by auto lemma privateKey_image_eq [simp]: - "(privateKey b A \ invKey ` publicKey c ` AS) = (b=c & A\AS)" + "(privateKey b A \ invKey ` publicKey c ` AS) = (b=c \ A\AS)" by auto lemma publicKey_notin_image_privateKey [simp]: "publicKey b A \ invKey ` publicKey c ` AS" @@ -239,7 +239,7 @@ apply (auto dest!: parts_cut simp add: used_Nil) done -lemma MPair_used_D: "\X,Y\ \ used H ==> X \ used H & Y \ used H" +lemma MPair_used_D: "\X,Y\ \ used H ==> X \ used H \ Y \ used H" by (drule used_parts_subset_parts, simp, blast) text\There was a similar theorem in Event.thy, so perhaps this one can @@ -358,7 +358,7 @@ subsection\Supply fresh nonces for possibility theorems\ text\In any trace, there is an upper bound N on the greatest nonce in use\ -lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \ used evs" +lemma Nonce_supply_lemma: "\N. \n. N\n \ Nonce n \ used evs" apply (induct_tac "evs") apply (rule_tac x = 0 in exI) apply (simp_all (no_asm_simp) add: used_Cons split: event.split) @@ -366,17 +366,17 @@ apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+ done -lemma Nonce_supply1: "EX N. Nonce N \ used evs" +lemma Nonce_supply1: "\N. Nonce N \ used evs" by (rule Nonce_supply_lemma [THEN exE], blast) -lemma Nonce_supply: "Nonce (@ N. Nonce N \ used evs) \ used evs" +lemma Nonce_supply: "Nonce (SOME N. Nonce N \ used evs) \ used evs" apply (rule Nonce_supply_lemma [THEN exE]) apply (rule someI, fast) done subsection\Specialized Rewriting for Theorems About @{term analz} and Image\ -lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H" +lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \ H" by blast lemma insert_Key_image: "insert (Key K) (Key`KK \ C) = Key ` (insert K KK) \ C" @@ -389,7 +389,7 @@ text\Lemma for the trivial direction of the if-and-only-if of the Session Key Compromise Theorem\ lemma analz_image_freshK_lemma: - "(Key K \ analz (Key`nE \ H)) --> (K \ nE | Key K \ analz H) ==> + "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) ==> (Key K \ analz (Key`nE \ H)) = (K \ nE | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Recur.thy --- a/src/HOL/Auth/Recur.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Recur.thy Thu Feb 15 13:04:36 2018 +0100 @@ -17,7 +17,7 @@ Perhaps the two session keys could be bundled into a single message. *) inductive_set (*Server's response to the nested message*) - respond :: "event list => (msg*msg*key)set" + respond :: "event list \ (msg*msg*key)set" for evs :: "event list" where One: "Key KAB \ used evs @@ -232,11 +232,11 @@ satisfying the inductive hypothesis.*) lemma resp_analz_image_freshK_lemma: "[| RB \ responses evs; - \K KK. KK \ - (range shrK) --> - (Key K \ analz (Key`KK Un H)) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ H)) = (K \ KK | Key K \ analz H) |] - ==> \K KK. KK \ - (range shrK) --> - (Key K \ analz (insert RB (Key`KK Un H))) = + ==> \K KK. KK \ - (range shrK) \ + (Key K \ analz (insert RB (Key`KK \ H))) = (K \ KK | Key K \ analz (insert RB H))" apply (erule responses.induct) apply (simp_all del: image_insert @@ -247,8 +247,8 @@ text\Version for the protocol. Proof is easy, thanks to the lemma.\ lemma raw_analz_image_freshK: "evs \ recur ==> - \K KK. KK \ - (range shrK) --> - (Key K \ analz (Key`KK Un (spies evs))) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ (spies evs))) = (K \ KK | Key K \ analz (spies evs))" apply (erule recur.induct) apply (drule_tac [4] RA2_analz_spies, @@ -261,8 +261,8 @@ (*Instance of the lemma with H replaced by (spies evs): [| RB \ responses evs; evs \ recur; |] - ==> KK \ - (range shrK) --> - Key K \ analz (insert RB (Key`KK Un spies evs)) = + ==> KK \ - (range shrK) \ + Key K \ analz (insert RB (Key`KK \ spies evs)) = (K \ KK | Key K \ analz (insert RB (spies evs))) *) lemmas resp_analz_image_freshK = @@ -302,7 +302,7 @@ "[| Hash \Key(shrK A), Agent A, B, NA, P\ \ parts (spies evs); Hash \Key(shrK A), Agent A, B',NA, P'\ \ parts (spies evs); evs \ recur; A \ bad |] - ==> B=B' & P=P'" + ==> B=B' \ P=P'" apply (erule rev_mp, erule rev_mp) apply (erule recur.induct, drule_tac [5] respond_imp_responses) @@ -322,7 +322,7 @@ lemma shrK_in_analz_respond [simp]: "[| RB \ responses evs; evs \ recur |] - ==> (Key (shrK B) \ analz (insert RB (spies evs))) = (B:bad)" + ==> (Key (shrK B) \ analz (insert RB (spies evs))) = (B\bad)" apply (erule responses.induct) apply (simp_all del: image_insert add: analz_image_freshK_simps resp_analz_image_freshK, auto) @@ -331,8 +331,8 @@ lemma resp_analz_insert_lemma: "[| Key K \ analz (insert RB H); - \K KK. KK \ - (range shrK) --> - (Key K \ analz (Key`KK Un H)) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ H)) = (K \ KK | Key K \ analz H); RB \ responses evs |] ==> (Key K \ parts{RB} | Key K \ analz H)" @@ -361,9 +361,9 @@ the quantifiers appear to be necessary.*) lemma unique_lemma [rule_format]: "(PB,RB,KXY) \ respond evs ==> - \A B N. Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB} --> - (\A' B' N'. Crypt (shrK A') \Key K, Agent B', N'\ \ parts {RB} --> - (A'=A & B'=B) | (A'=B & B'=A))" + \A B N. Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB} \ + (\A' B' N'. Crypt (shrK A') \Key K, Agent B', N'\ \ parts {RB} \ + (A'=A \ B'=B) | (A'=B \ B'=A))" apply (erule respond.induct) apply (simp_all add: all_conj_distrib) apply (blast dest: respond_certificate) @@ -373,7 +373,7 @@ "[| Crypt (shrK A) \Key K, Agent B, N\ \ parts {RB}; Crypt (shrK A') \Key K, Agent B', N'\ \ parts {RB}; (PB,RB,KXY) \ respond evs |] - ==> (A'=A & B'=B) | (A'=B & B'=A)" + ==> (A'=A \ B'=B) | (A'=B \ B'=A)" by (rule unique_lemma, auto) @@ -383,8 +383,8 @@ lemma respond_Spy_not_see_session_key [rule_format]: "[| (PB,RB,KAB) \ respond evs; evs \ recur |] - ==> \A A' N. A \ bad & A' \ bad --> - Crypt (shrK A) \Key K, Agent A', N\ \ parts{RB} --> + ==> \A A' N. A \ bad \ A' \ bad \ + Crypt (shrK A) \Key K, Agent A', N\ \ parts{RB} \ Key K \ analz (insert RB (spies evs))" apply (erule respond.induct) apply (frule_tac [2] respond_imp_responses) @@ -464,7 +464,7 @@ lemma Cert_imp_Server_msg: "[| Crypt (shrK A) Y \ parts (spies evs); A \ bad; evs \ recur |] - ==> \C RC. Says Server C RC \ set evs & + ==> \C RC. Says Server C RC \ set evs \ Crypt (shrK A) Y \ parts {RC}" apply (erule rev_mp, erule recur.induct, simp_all) txt\Fake\ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Shared.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,7 +12,7 @@ begin consts - shrK :: "agent => key" (*symmetric keys*) + shrK :: "agent \ key" (*symmetric keys*) specification (shrK) inj_shrK: "inj shrK" @@ -76,13 +76,13 @@ subsection\Function "knows"\ (*Spy sees shared keys of agents!*) -lemma Spy_knows_Spy_bad [intro!]: "A: bad ==> Key (shrK A) \ knows Spy evs" +lemma Spy_knows_Spy_bad [intro!]: "A \ bad \ Key (shrK A) \ knows Spy evs" apply (induct_tac "evs") apply (simp_all (no_asm_simp) add: imageI knows_Cons split: event.split) done (*For case analysis on whether or not an agent is compromised*) -lemma Crypt_Spy_analz_bad: "[| Crypt (shrK A) X \ analz (knows Spy evs); A: bad |] +lemma Crypt_Spy_analz_bad: "[| Crypt (shrK A) X \ analz (knows Spy evs); A \ bad |] ==> X \ analz (knows Spy evs)" by (metis Spy_knows_Spy_bad analz.Inj analz_Decrypt') @@ -120,7 +120,7 @@ subsection\Supply fresh nonces for possibility theorems.\ (*In any trace, there is an upper bound N on the greatest nonce in use.*) -lemma Nonce_supply_lemma: "\N. ALL n. N<=n --> Nonce n \ used evs" +lemma Nonce_supply_lemma: "\N. \n. N \ n \ Nonce n \ used evs" apply (induct_tac "evs") apply (rule_tac x = 0 in exI) apply (simp_all (no_asm_simp) add: used_Cons split: event.split) @@ -130,14 +130,14 @@ lemma Nonce_supply1: "\N. Nonce N \ used evs" by (metis Nonce_supply_lemma order_eq_iff) -lemma Nonce_supply2: "\N N'. Nonce N \ used evs & Nonce N' \ used evs' & N \ N'" +lemma Nonce_supply2: "\N N'. Nonce N \ used evs \ Nonce N' \ used evs' \ N \ N'" apply (cut_tac evs = evs in Nonce_supply_lemma) apply (cut_tac evs = "evs'" in Nonce_supply_lemma, clarify) apply (metis Suc_n_not_le_n nat_le_linear) done -lemma Nonce_supply3: "\N N' N''. Nonce N \ used evs & Nonce N' \ used evs' & - Nonce N'' \ used evs'' & N \ N' & N' \ N'' & N \ N''" +lemma Nonce_supply3: "\N N' N''. Nonce N \ used evs \ Nonce N' \ used evs' \ + Nonce N'' \ used evs'' \ N \ N' \ N' \ N'' \ N \ N''" apply (cut_tac evs = evs in Nonce_supply_lemma) apply (cut_tac evs = "evs'" in Nonce_supply_lemma) apply (cut_tac evs = "evs''" in Nonce_supply_lemma, clarify) @@ -147,13 +147,13 @@ apply (simp (no_asm_simp) add: less_not_refl3 le_add1 le_add2 less_Suc_eq_le) done -lemma Nonce_supply: "Nonce (@ N. Nonce N \ used evs) \ used evs" +lemma Nonce_supply: "Nonce (SOME N. Nonce N \ used evs) \ used evs" apply (rule Nonce_supply_lemma [THEN exE]) apply (rule someI, blast) done text\Unlike the corresponding property of nonces, we cannot prove - @{term "finite KK ==> \K. K \ KK & Key K \ used evs"}. + @{term "finite KK ==> \K. K \ KK \ Key K \ used evs"}. We have infinitely many agents and there is nothing to stop their long-term keys from exhausting all the natural numbers. Instead, possibility theorems must assume the existence of a few keys.\ @@ -161,7 +161,7 @@ subsection\Specialized Rewriting for Theorems About @{term analz} and Image\ -lemma subset_Compl_range: "A <= - (range shrK) ==> shrK x \ A" +lemma subset_Compl_range: "A \ - (range shrK) \ shrK x \ A" by blast lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \ H" @@ -184,7 +184,7 @@ (*Lemma for the trivial direction of the if-and-only-if*) lemma analz_image_freshK_lemma: - "(Key K \ analz (Key`nE \ H)) --> (K \ nE | Key K \ analz H) ==> + "(Key K \ analz (Key`nE \ H)) \ (K \ nE | Key K \ analz H) ==> (Key K \ analz (Key`nE \ H)) = (K \ nE | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) @@ -250,7 +250,7 @@ Scan.succeed (fn ctxt => SIMPLE_METHOD (Shared.basic_possibility_tac ctxt))\ "for proving possibility theorems" -lemma knows_subset_knows_Cons: "knows A evs <= knows A (e # evs)" +lemma knows_subset_knows_Cons: "knows A evs \ knows A (e # evs)" by (cases e) (auto simp: knows_Cons) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Thu Feb 15 13:04:36 2018 +0100 @@ -299,7 +299,7 @@ apply blast done -lemma Nonce_supply: "Nonce (@ N. Nonce N \ used evs) \ used evs" +lemma Nonce_supply: "Nonce (SOME N. Nonce N \ used evs) \ used evs" apply (rule finite.emptyI [THEN Nonce_supply_ax, THEN exE]) apply (rule someI, blast) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/TLS.thy --- a/src/HOL/Auth/TLS.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/TLS.thy Thu Feb 15 13:04:36 2018 +0100 @@ -42,7 +42,7 @@ theory TLS imports Public "HOL-Library.Nat_Bijection" begin -definition certificate :: "[agent,key] => msg" where +definition certificate :: "[agent,key] \ msg" where "certificate A KA == Crypt (priSK Server) \Agent A, Key KA\" text\TLS apparently does not require separate keypairs for encryption and @@ -53,27 +53,27 @@ consts (*Pseudo-random function of Section 5*) - PRF :: "nat*nat*nat => nat" + PRF :: "nat*nat*nat \ nat" (*Client, server write keys are generated uniformly by function sessionK to avoid duplicating their properties. They are distinguished by a tag (not a bool, to avoid the peculiarities of if-and-only-if). Session keys implicitly include MAC secrets.*) - sessionK :: "(nat*nat*nat) * role => key" + sessionK :: "(nat*nat*nat) * role \ key" abbreviation - clientK :: "nat*nat*nat => key" where + clientK :: "nat*nat*nat \ key" where "clientK X == sessionK(X, ClientRole)" abbreviation - serverK :: "nat*nat*nat => key" where + serverK :: "nat*nat*nat \ key" where "serverK X == sessionK(X, ServerRole)" specification (PRF) inj_PRF: "inj PRF" \ \the pseudo-random function is collision-free\ - apply (rule exI [of _ "%(x,y,z). prod_encode(x, prod_encode(y,z))"]) + apply (rule exI [of _ "\(x,y,z). prod_encode(x, prod_encode(y,z))"]) apply (simp add: inj_on_def prod_encode_eq) done @@ -81,7 +81,7 @@ inj_sessionK: "inj sessionK" \ \sessionK is collision-free; also, no clientK clashes with any serverK.\ apply (rule exI [of _ - "%((x,y,z), r). prod_encode(case_role 0 1 r, + "\((x,y,z), r). prod_encode(case_role 0 1 r, prod_encode(x, prod_encode(y,z)))"]) apply (simp add: inj_on_def prod_encode_eq split: role.split) done @@ -108,7 +108,7 @@ | SpyKeys: \ \The spy may apply @{term PRF} and @{term sessionK} to available nonces\ "[| evsSK \ tls; - {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |] + {Nonce NA, Nonce NB, Nonce M} \ analz (spies evsSK) |] ==> Notes Spy \ Nonce (PRF(M,NA,NB)), Key (sessionK((NA,NB,M),role))\ # evsSK \ tls" @@ -240,7 +240,7 @@ \ \If A recalls the \SESSION_ID\, then she sends a FINISHED message using the new nonces and stored MASTER SECRET.\ "[| evsCR \ tls; - Says A B \Agent A, Nonce NA, Number SID, Number PA\: set evsCR; + Says A B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsCR; Says B' A \Nonce NB, Number SID, Number PB\ \ set evsCR; Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evsCR |] ==> Says A B (Crypt (clientK(NA,NB,M)) @@ -253,7 +253,7 @@ \ \Resumption (7.3): If B finds the \SESSION_ID\ then he can send a FINISHED message using the recovered MASTER SECRET\ "[| evsSR \ tls; - Says A' B \Agent A, Nonce NA, Number SID, Number PA\: set evsSR; + Says A' B \Agent A, Nonce NA, Number SID, Number PA\ \ set evsSR; Says B A \Nonce NB, Number SID, Number PB\ \ set evsSR; Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evsSR |] ==> Says B A (Crypt (serverK(NA,NB,M)) @@ -331,7 +331,7 @@ text\Possibility property ending with ClientAccepts.\ -lemma "[| \evs. (@ N. Nonce N \ used evs) \ range PRF; A \ B |] +lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] ==> \SID M. \evs \ tls. Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evs" apply (intro exI bexI) @@ -344,7 +344,7 @@ text\And one for ServerAccepts. Either FINISHED message may come first.\ -lemma "[| \evs. (@ N. Nonce N \ used evs) \ range PRF; A \ B |] +lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] ==> \SID NA PA NB PB M. \evs \ tls. Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evs" apply (intro exI bexI) @@ -357,7 +357,7 @@ text\Another one, for CertVerify (which is optional)\ -lemma "[| \evs. (@ N. Nonce N \ used evs) \ range PRF; A \ B |] +lemma "[| \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] ==> \NB PMS. \evs \ tls. Says A B (Crypt (priK A) (Hash\Nonce NB, Agent B, Nonce PMS\)) \ set evs" @@ -374,13 +374,13 @@ lemma "[| evs0 \ tls; Notes A \Number SID, Agent A, Agent B, Nonce M\ \ set evs0; Notes B \Number SID, Agent A, Agent B, Nonce M\ \ set evs0; - \evs. (@ N. Nonce N \ used evs) \ range PRF; + \evs. (SOME N. Nonce N \ used evs) \ range PRF; A \ B |] ==> \NA PA NB PB X. \evs \ tls. X = Hash\Number SID, Nonce M, Nonce NA, Number PA, Agent A, - Nonce NB, Number PB, Agent B\ & - Says A B (Crypt (clientK(NA,NB,M)) X) \ set evs & + Nonce NB, Number PB, Agent B\ \ + Says A B (Crypt (clientK(NA,NB,M)) X) \ set evs \ Says B A (Crypt (serverK(NA,NB,M)) X) \ set evs" apply (intro exI bexI) apply (rule_tac [2] tls.ClientHello @@ -548,7 +548,7 @@ "[| Notes A \Agent B, Nonce PMS\ \ set evs; Notes A' \Agent B', Nonce PMS\ \ set evs; evs \ tls |] - ==> A=A' & B=B'" + ==> A=A' \ B=B'" apply (erule rev_mp, erule rev_mp) apply (erule tls.induct, force, simp_all) txt\ClientKeyExch\ @@ -562,7 +562,7 @@ No collection of keys can help the spy get new private keys.\ lemma analz_image_priK [rule_format]: "evs \ tls - ==> \KK. (Key(priK B) \ analz (Key`KK Un (spies evs))) = + ==> \KK. (Key(priK B) \ analz (Key`KK \ (spies evs))) = (priK B \ KK | B \ bad)" apply (erule tls.induct) apply (simp_all (no_asm_simp) @@ -576,25 +576,25 @@ text\slightly speeds up the big simplification below\ lemma range_sessionkeys_not_priK: - "KK <= range sessionK ==> priK B \ KK" + "KK \ range sessionK \ priK B \ KK" by blast text\Lemma for the trivial direction of the if-and-only-if\ lemma analz_image_keys_lemma: - "(X \ analz (G Un H)) --> (X \ analz H) ==> - (X \ analz (G Un H)) = (X \ analz H)" + "(X \ analz (G \ H)) \ (X \ analz H) ==> + (X \ analz (G \ H)) = (X \ analz H)" by (blast intro: analz_mono [THEN subsetD]) (** Strangely, the following version doesn't work: -\Z. (Nonce N \ analz (Key`(sessionK`Z) Un (spies evs))) = +\Z. (Nonce N \ analz (Key`(sessionK`Z) \ (spies evs))) = (Nonce N \ analz (spies evs))" **) lemma analz_image_keys [rule_format]: "evs \ tls ==> - \KK. KK <= range sessionK --> - (Nonce N \ analz (Key`KK Un (spies evs))) = + \KK. KK \ range sessionK \ + (Nonce N \ analz (Key`KK \ (spies evs))) = (Nonce N \ analz (spies evs))" apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (safe del: iffI) @@ -631,7 +631,7 @@ "[| Nonce PMS \ parts (spies evs); K = sessionK((Na, Nb, PRF(PMS,NA,NB)), role); evs \ tls |] - ==> Key K \ parts (spies evs) & (\Y. Crypt K Y \ parts (spies evs))" + ==> Key K \ parts (spies evs) \ (\Y. Crypt K Y \ parts (spies evs))" apply (erule rev_mp, erule ssubst) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) @@ -809,9 +809,9 @@ Nonce Nb, Number PB, Agent B\); M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad |] - ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs --> - Notes A \Agent B, Nonce PMS\ \ set evs --> - X \ parts (spies evs) --> Says B A X \ set evs" + ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ + Notes A \Agent B, Nonce PMS\ \ set evs \ + X \ parts (spies evs) \ Says B A X \ set evs" apply (erule ssubst)+ apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) apply (force, simp_all (no_asm_simp)) @@ -828,9 +828,9 @@ to bind A's identity with PMS, then we could replace A' by A below.\ lemma TrustServerMsg [rule_format]: "[| M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad |] - ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs --> - Notes A \Agent B, Nonce PMS\ \ set evs --> - Crypt (serverK(Na,Nb,M)) Y \ parts (spies evs) --> + ==> Says B Spy (Key(serverK(Na,Nb,M))) \ set evs \ + Notes A \Agent B, Nonce PMS\ \ set evs \ + Crypt (serverK(Na,Nb,M)) Y \ parts (spies evs) \ (\A'. Says B A' (Crypt (serverK(Na,Nb,M)) Y) \ set evs)" apply (erule ssubst) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) @@ -854,9 +854,9 @@ lemma TrustClientMsg [rule_format]: "[| M = PRF(PMS,NA,NB); evs \ tls; A \ bad; B \ bad |] - ==> Says A Spy (Key(clientK(Na,Nb,M))) \ set evs --> - Notes A \Agent B, Nonce PMS\ \ set evs --> - Crypt (clientK(Na,Nb,M)) Y \ parts (spies evs) --> + ==> Says A Spy (Key(clientK(Na,Nb,M))) \ set evs \ + Notes A \Agent B, Nonce PMS\ \ set evs \ + Crypt (clientK(Na,Nb,M)) Y \ parts (spies evs) \ Says A B (Crypt (clientK(Na,Nb,M)) Y) \ set evs" apply (erule ssubst) apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/WooLam.thy --- a/src/HOL/Auth/WooLam.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/WooLam.thy Thu Feb 15 13:04:36 2018 +0100 @@ -142,7 +142,7 @@ But A may have sent the nonce to some other agent and it could have reached the Server via the Spy.*) lemma B_trusts_WL5: - "[| Says S B (Crypt (shrK B) \Agent A, Nonce NB\): set evs; + "[| Says S B (Crypt (shrK B) \Agent A, Nonce NB\) \ set evs; A \ bad; B \ bad; evs \ woolam |] ==> \B. Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" by (blast dest!: NB_Crypt_imp_Server_msg) @@ -157,9 +157,9 @@ (**CANNOT be proved because A doesn't know where challenges come from...*) lemma "[| A \ bad; B \ Spy; evs \ woolam |] - ==> Crypt (shrK A) (Nonce NB) \ parts (spies evs) & + ==> Crypt (shrK A) (Nonce NB) \ parts (spies evs) \ Says B A (Nonce NB) \ set evs - --> Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" + \ Says A B (Crypt (shrK A) (Nonce NB)) \ set evs" apply (erule rev_mp, erule woolam.induct, force, simp_all, blast, auto) oops diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Yahalom.thy --- a/src/HOL/Auth/Yahalom.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Yahalom.thy Thu Feb 15 13:04:36 2018 +0100 @@ -73,7 +73,7 @@ \ Notes Spy \Nonce NA, Nonce NB, Key K\ # evso \ yahalom" -definition KeyWithNonce :: "[key, nat, event list] => bool" where +definition KeyWithNonce :: "[key, nat, event list] \ bool" where "KeyWithNonce K NB evs == \A B na X. Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB\, X\ @@ -191,8 +191,8 @@ lemma analz_image_freshK [rule_format]: "evs \ yahalom \ - \K KK. KK <= - (range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule yahalom.induct, drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast) @@ -214,7 +214,7 @@ Says Server A' \Crypt (shrK A') \Agent B', Key K, na', nb'\, X'\ \ set evs; evs \ yahalom\ - \ A=A' & B=B' & na=na' & nb=nb'" + \ A=A' \ B=B' \ na=na' \ nb=nb'" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, simp_all) txt\YM3, by freshness, and YM4\ @@ -228,8 +228,8 @@ \ Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ - \ set evs --> - Notes Spy \na, nb, Key K\ \ set evs --> + \ set evs \ + Notes Spy \na, nb, Key K\ \ set evs \ Key K \ analz (knows Spy evs)" apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy) @@ -336,7 +336,7 @@ lemma KeyWithNonce_Says [simp]: "KeyWithNonce K NB (Says S A X # evs) = - (Server = S & + (Server = S \ (\B n X'. X = \Crypt (shrK A) \Agent B, Key K, n, Nonce NB\, X'\) | KeyWithNonce K NB evs)" by (simp add: KeyWithNonce_def, blast) @@ -353,7 +353,7 @@ text\A fresh key cannot be associated with any nonce (with respect to a given trace).\ lemma fresh_not_KeyWithNonce: - "Key K \ used evs \ ~ KeyWithNonce K NB evs" + "Key K \ used evs \ \ KeyWithNonce K NB evs" by (unfold KeyWithNonce_def, blast) text\The Server message associates K with NB' and therefore not with any @@ -362,7 +362,7 @@ "\Says Server A \Crypt (shrK A) \Agent B, Key K, na, Nonce NB'\, X\ \ set evs; NB \ NB'; evs \ yahalom\ - \ ~ KeyWithNonce K NB evs" + \ \ KeyWithNonce K NB evs" by (unfold KeyWithNonce_def, blast dest: unique_session_keys) @@ -374,15 +374,15 @@ text\As with \analz_image_freshK\, we take some pains to express the property as a logical equivalence so that the simplifier can apply it.\ lemma Nonce_secrecy_lemma: - "P --> (X \ analz (G Un H)) --> (X \ analz H) \ - P --> (X \ analz (G Un H)) = (X \ analz H)" + "P \ (X \ analz (G \ H)) \ (X \ analz H) \ + P \ (X \ analz (G \ H)) = (X \ analz H)" by (blast intro: analz_mono [THEN subsetD]) lemma Nonce_secrecy: "evs \ yahalom \ - (\KK. KK <= - (range shrK) --> - (\K \ KK. K \ symKeys --> ~ KeyWithNonce K NB evs) --> - (Nonce NB \ analz (Key`KK Un (knows Spy evs))) = + (\KK. KK \ - (range shrK) \ + (\K \ KK. K \ symKeys \ \ KeyWithNonce K NB evs) \ + (Nonce NB \ analz (Key`KK \ (knows Spy evs))) = (Nonce NB \ analz (knows Spy evs)))" apply (erule yahalom.induct, frule_tac [7] YM4_analz_knows_Spy) @@ -394,7 +394,7 @@ imp_disj_not1 (*Moves NBa\NB to the front*) Says_Server_KeyWithNonce) txt\For Oops, simplification proves @{prop "NBa\NB"}. By - @{term Says_Server_KeyWithNonce}, we get @{prop "~ KeyWithNonce K NB + @{term Says_Server_KeyWithNonce}, we get @{prop "\ KeyWithNonce K NB evs"}; then simplification can apply the induction hypothesis with @{term "KK = {K}"}.\ subgoal \ \Fake\ by spy_analz @@ -427,7 +427,7 @@ "\Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); Crypt (shrK B') \Agent A', Nonce NA', nb\ \ parts (knows Spy evs); evs \ yahalom; B \ bad; B' \ bad\ - \ NA' = NA & A' = A & B' = B" + \ NA' = NA \ A' = A \ B' = B" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy, simp_all) @@ -444,7 +444,7 @@ Gets S' \X', Crypt (shrK B') \Agent A', Nonce NA', nb\\ \ set evs; nb \ analz (knows Spy evs); evs \ yahalom\ - \ NA' = NA & A' = A & B' = B" + \ NA' = NA \ A' = A \ B' = B" by (blast dest!: Gets_imp_Says Crypt_Spy_analz_bad dest: Says_imp_spies unique_NB parts.Inj analz.Inj) @@ -548,7 +548,7 @@ lemma B_Said_YM2 [rule_format]: "\Crypt (shrK B) \Agent A, Nonce NA, nb\ \ parts (knows Spy evs); evs \ yahalom\ - \ B \ bad --> + \ B \ bad \ Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs" apply (erule rev_mp, erule yahalom.induct, force, @@ -561,7 +561,7 @@ lemma YM3_auth_B_to_A_lemma: "\Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, nb\, X\ \ set evs; evs \ yahalom\ - \ B \ bad --> + \ B \ bad \ Says B Server \Agent B, Crypt (shrK B) \Agent A, Nonce NA, nb\\ \ set evs" apply (erule rev_mp, erule yahalom.induct, simp_all) @@ -588,10 +588,10 @@ NB matters for freshness.\ theorem A_Said_YM3_lemma [rule_format]: "evs \ yahalom - \ Key K \ analz (knows Spy evs) --> - Crypt K (Nonce NB) \ parts (knows Spy evs) --> - Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs) --> - B \ bad --> + \ Key K \ analz (knows Spy evs) \ + Crypt K (Nonce NB) \ parts (knows Spy evs) \ + Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs) \ + B \ bad \ (\X. Says A B \X, Crypt K (Nonce NB)\ \ set evs)" apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Yahalom2.thy --- a/src/HOL/Auth/Yahalom2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Yahalom2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -172,8 +172,8 @@ lemma analz_image_freshK [rule_format]: "evs \ yahalom \ - \K KK. KK <= - (range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" apply (erule yahalom.induct) apply (frule_tac [8] Says_Server_message_form) @@ -194,7 +194,7 @@ Says Server A' \nb', Crypt (shrK A') \Agent B', Key K, na'\, X'\ \ set evs; evs \ yahalom\ - \ A=A' & B=B' & na=na' & nb=nb'" + \ A=A' \ B=B' \ na=na' \ nb=nb'" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, simp_all) txt\YM3, by freshness\ @@ -209,8 +209,8 @@ \ Says Server A \nb, Crypt (shrK A) \Agent B, Key K, na\, Crypt (shrK B) \Agent A, Agent B, Key K, nb\\ - \ set evs --> - Notes Spy \na, nb, Key K\ \ set evs --> + \ set evs \ + Notes Spy \na, nb, Key K\ \ set evs \ Key K \ analz (knows Spy evs)" apply (erule yahalom.induct, force, frule_tac [7] Says_Server_message_form, drule_tac [6] YM4_analz_knows_Spy) @@ -384,18 +384,18 @@ "\Crypt (shrK A) \Agent B, Key K, na\ \ analz (spies evs); Crypt (shrK A') \Agent B', Key K, na'\ \ analz (spies evs); Key K \ analz (knows Spy evs); evs \ yahalom\ - \ A=A' & B=B'" + \ A=A' \ B=B'" by (blast dest!: A_trusts_YM3 dest: unique_session_keys Crypt_Spy_analz_bad) lemma Auth_A_to_B_lemma [rule_format]: "evs \ yahalom - \ Key K \ analz (knows Spy evs) --> - K \ symKeys --> - Crypt K (Nonce NB) \ parts (knows Spy evs) --> + \ Key K \ analz (knows Spy evs) \ + K \ symKeys \ + Crypt K (Nonce NB) \ parts (knows Spy evs) \ Crypt (shrK B) \Agent A, Agent B, Key K, Nonce NB\ - \ parts (knows Spy evs) --> - B \ bad --> + \ parts (knows Spy evs) \ + B \ bad \ (\X. Says A B \X, Crypt K (Nonce NB)\ \ set evs)" apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Auth/Yahalom_Bad.thy --- a/src/HOL/Auth/Yahalom_Bad.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Auth/Yahalom_Bad.thy Thu Feb 15 13:04:36 2018 +0100 @@ -152,8 +152,8 @@ lemma analz_image_freshK [rule_format]: "evs \ yahalom ==> - \K KK. KK <= - (range shrK) --> - (Key K \ analz (Key`KK Un (knows Spy evs))) = + \K KK. KK \ - (range shrK) \ + (Key K \ analz (Key`KK \ (knows Spy evs))) = (K \ KK | Key K \ analz (knows Spy evs))" by (erule yahalom.induct, drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast) @@ -172,7 +172,7 @@ Says Server A' \Crypt (shrK A') \Agent B', Key K, na', nb'\, X'\ \ set evs; evs \ yahalom |] - ==> A=A' & B=B' & na=na' & nb=nb'" + ==> A=A' \ B=B' \ na=na' \ nb=nb'" apply (erule rev_mp, erule rev_mp) apply (erule yahalom.induct, simp_all) txt\YM3, by freshness, and YM4\ @@ -186,7 +186,7 @@ ==> Says Server A \Crypt (shrK A) \Agent B, Key K, na, nb\, Crypt (shrK B) \Agent A, Key K\\ - \ set evs --> + \ set evs \ Key K \ analz (knows Spy evs)" apply (erule yahalom.induct, force, drule_tac [6] YM4_analz_knows_Spy) apply (simp_all add: pushes analz_insert_eq analz_insert_freshK, spy_analz) (*Fake*) @@ -260,7 +260,7 @@ the secrecy of NB.\ lemma B_trusts_YM4_newK [rule_format]: "[|Key K \ analz (knows Spy evs); evs \ yahalom|] - ==> Crypt K (Nonce NB) \ parts (knows Spy evs) --> + ==> Crypt K (Nonce NB) \ parts (knows Spy evs) \ (\A B NA. Says Server A \Crypt (shrK A) \Agent B, Key K, Nonce NA, Nonce NB\, @@ -323,10 +323,10 @@ NB matters for freshness.\ lemma A_Said_YM3_lemma [rule_format]: "evs \ yahalom - ==> Key K \ analz (knows Spy evs) --> - Crypt K (Nonce NB) \ parts (knows Spy evs) --> - Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs) --> - B \ bad --> + ==> Key K \ analz (knows Spy evs) \ + Crypt K (Nonce NB) \ parts (knows Spy evs) \ + Crypt (shrK B) \Agent A, Key K\ \ parts (knows Spy evs) \ + B \ bad \ (\X. Says A B \X, Crypt K (Nonce NB)\ \ set evs)" apply (erule yahalom.induct, force, frule_tac [6] YM4_parts_knows_Spy) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Cardinal_Arithmetic.thy --- a/src/HOL/BNF_Cardinal_Arithmetic.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Cardinal_Arithmetic.thy Thu Feb 15 13:04:36 2018 +0100 @@ -689,7 +689,7 @@ lemma cardSuc_UNION_Cinfinite: assumes "Cinfinite r" "relChain (cardSuc r) As" "B \ (UN i : Field (cardSuc r). As i)" "|B| <=o r" - shows "EX i : Field (cardSuc r). B \ As i" + shows "\i \ Field (cardSuc r). B \ As i" using cardSuc_UNION assms unfolding cinfinite_def by blast end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Cardinal_Order_Relation.thy --- a/src/HOL/BNF_Cardinal_Order_Relation.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Cardinal_Order_Relation.thy Thu Feb 15 13:04:36 2018 +0100 @@ -290,7 +290,7 @@ using ordLeq_iff_ordLess_Restr card_of_Well_order by blast lemma card_of_underS: -assumes r: "Card_order r" and a: "a : Field r" +assumes r: "Card_order r" and a: "a \ Field r" shows "|underS r a| finite (Field r)" -and a: "a : Field r" -shows "EX b : Field r. a \ b \ (a,b) : r" +and a: "a \ Field r" +shows "\b \ Field r. a \ b \ (a,b) \ r" proof- have "Field r \ under r a" using assms Card_order_infinite_not_under by blast @@ -891,7 +891,7 @@ moreover have ba: "b \ a" using 1 r unfolding card_order_on_def well_order_on_def linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto - ultimately have "(a,b) : r" + ultimately have "(a,b) \ r" using a r unfolding card_order_on_def well_order_on_def linear_order_on_def total_on_def by blast thus ?thesis using 1 ba by auto @@ -1506,7 +1506,7 @@ definition cofinal where "cofinal A r \ - ALL a : Field r. EX b : A. a \ b \ (a,b) : r" + \a \ Field r. \b \ A. a \ b \ (a,b) \ r" definition regularCard where "regularCard r \ @@ -1521,21 +1521,21 @@ and As: "relChain r As" and Bsub: "B \ (UN i : Field r. As i)" and cardB: "|B| As i" +shows "\i \ Field r. B \ As i" proof- let ?phi = "\b j. j \ Field r \ b \ As j" have "\b\B. \j. ?phi b j" using Bsub by blast - then obtain f where f: "!! b. b : B \ ?phi b (f b)" + then obtain f where f: "\b. b \ B \ ?phi b (f b)" using bchoice[of B ?phi] by blast let ?K = "f ` B" {assume 1: "\i. i \ Field r \ \ B \ As i" have 2: "cofinal ?K r" unfolding cofinal_def proof auto - fix i assume i: "i : Field r" - with 1 obtain b where b: "b : B \ b \ As i" by blast + fix i assume i: "i \ Field r" + with 1 obtain b where b: "b \ B \ b \ As i" by blast hence "i \ f b \ \ (f b,i) \ r" using As f unfolding relChain_def by auto - hence "i \ f b \ (i, f b) : r" using r + hence "i \ f b \ (i, f b) \ r" using r unfolding card_order_on_def well_order_on_def linear_order_on_def total_on_def using i f b by auto with b show "\b\B. i \ f b \ (i, f b) \ r" by blast @@ -1576,11 +1576,11 @@ hence "|K| <=o r" using r' card_of_Card_order[of K] by blast hence "|K| \o |?J|" using rJ ordLeq_ordIso_trans by blast moreover - {have "ALL j : K. |underS ?r' j| j\K. |underS ?r' j| o r" + hence "\j\K. |underS ?r' j| \o r" using r' card_of_Card_order by blast - hence "ALL j : K. |underS ?r' j| \o |?J|" + hence "\j\K. |underS ?r' j| \o |?J|" using rJ ordLeq_ordIso_trans by blast } ultimately have "|?L| \o |?J|" @@ -1608,7 +1608,7 @@ and As: "relChain (cardSuc r) As" and Bsub: "B \ (UN i : Field (cardSuc r). As i)" and cardB: "|B| <=o r" -shows "EX i : Field (cardSuc r). B \ As i" +shows "\i \ Field (cardSuc r). B \ As i" proof- let ?r' = "cardSuc r" have "Card_order ?r' \ |B| P = Q" by blast -lemma Grp_fst_snd: "(Grp (Collect (case_prod R)) fst)^--1 OO Grp (Collect (case_prod R)) snd = R" +lemma Grp_fst_snd: "(Grp (Collect (case_prod R)) fst)\\ OO Grp (Collect (case_prod R)) snd = R" unfolding Grp_def fun_eq_iff relcompp.simps by auto -lemma OO_Grp_cong: "A = B \ (Grp A f)^--1 OO Grp A g = (Grp B f)^--1 OO Grp B g" +lemma OO_Grp_cong: "A = B \ (Grp A f)\\ OO Grp A g = (Grp B f)\\ OO Grp B g" by (rule arg_cong) lemma vimage2p_relcompp_mono: "R OO S \ T \ @@ -145,7 +145,7 @@ qed blast lemma vimage2p_relcompp_converse: - "vimage2p f g (R^--1 OO S) = (vimage2p Rep f R)^--1 OO vimage2p Rep g S" + "vimage2p f g (R\\ OO S) = (vimage2p Rep f R)\\ OO vimage2p Rep g S" unfolding vimage2p_def relcompp.simps conversep.simps fun_eq_iff image_def by (auto simp: type_copy_ex_RepI) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Def.thy --- a/src/HOL/BNF_Def.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Def.thy Thu Feb 15 13:04:36 2018 +0100 @@ -105,16 +105,16 @@ lemma eq_alt: "(=) = Grp UNIV id" unfolding Grp_def by auto -lemma leq_conversepI: "R = (=) \ R \ R^--1" +lemma leq_conversepI: "R = (=) \ R \ R\\" by auto lemma leq_OOI: "R = (=) \ R \ R OO R" by auto -lemma OO_Grp_alt: "(Grp A f)^--1 OO Grp A g = (\x y. \z. z \ A \ f z = x \ g z = y)" +lemma OO_Grp_alt: "(Grp A f)\\ OO Grp A g = (\x y. \z. z \ A \ f z = x \ g z = y)" unfolding Grp_def by auto -lemma Grp_UNIV_id: "f = id \ (Grp UNIV f)^--1 OO Grp UNIV f = Grp UNIV f" +lemma Grp_UNIV_id: "f = id \ (Grp UNIV f)\\ OO Grp UNIV f = Grp UNIV f" unfolding Grp_def by auto lemma Grp_UNIV_idI: "x = y \ Grp UNIV id x y" @@ -171,7 +171,7 @@ lemma fst_snd_flip: "fst xy = (snd \ (%(x, y). (y, x))) xy" by (simp split: prod.split) -lemma flip_pred: "A \ Collect (case_prod (R ^--1)) \ (%(x, y). (y, x)) ` A \ Collect (case_prod R)" +lemma flip_pred: "A \ Collect (case_prod (R \\)) \ (%(x, y). (y, x)) ` A \ Collect (case_prod R)" by auto lemma predicate2_eqD: "A = B \ A a b \ B a b" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Greatest_Fixpoint.thy --- a/src/HOL/BNF_Greatest_Fixpoint.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Greatest_Fixpoint.thy Thu Feb 15 13:04:36 2018 +0100 @@ -31,7 +31,7 @@ lemma neq_eq_eq_contradict: "\t \ u; s = t; s = u\ \ P" by fast -lemma converse_Times: "(A \ B) ^-1 = B \ A" +lemma converse_Times: "(A \ B)\ = B \ A" by fast lemma equiv_proj: @@ -56,7 +56,7 @@ lemma IdD: "(a, b) \ Id \ a = b" by auto -lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)" +lemma image2_Gr: "image2 A f g = (Gr A f)\ O (Gr A g)" unfolding image2_def Gr_def by auto lemma GrD1: "(x, fx) \ Gr A f \ x \ A" @@ -99,10 +99,10 @@ "relInvImage A R f \ {(a1, a2) | a1 a2. a1 \ A \ a2 \ A \ (f a1, f a2) \ R}" lemma relImage_Gr: - "\R \ A \ A\ \ relImage R f = (Gr A f)^-1 O R O Gr A f" + "\R \ A \ A\ \ relImage R f = (Gr A f)\ O R O Gr A f" unfolding relImage_def Gr_def relcomp_def by auto -lemma relInvImage_Gr: "\R \ B \ B\ \ relInvImage A R f = Gr A f O R O (Gr A f)^-1" +lemma relInvImage_Gr: "\R \ B \ B\ \ relInvImage A R f = Gr A f O R O (Gr A f)\" unfolding Gr_def relcomp_def image_def relInvImage_def by auto lemma relImage_mono: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Wellorder_Constructions.thy --- a/src/HOL/BNF_Wellorder_Constructions.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Wellorder_Constructions.thy Thu Feb 15 13:04:36 2018 +0100 @@ -838,7 +838,7 @@ where "ord_to_filter r0 r \ (SOME f. embed r r0 f) ` (Field r)" lemma ord_to_filter_compat: -"compat (ordLess Int (ordLess^-1``{r0} \ ordLess^-1``{r0})) +"compat (ordLess Int (ordLess\``{r0} \ ordLess\``{r0})) (ofilterIncl r0) (ord_to_filter r0)" proof(unfold compat_def ord_to_filter_def, clarify) @@ -859,7 +859,7 @@ {fix r0 :: "('a \ 'a) set" (* need to annotate here!*) let ?ordLess = "ordLess::('d rel * 'd rel) set" - let ?R = "?ordLess Int (?ordLess^-1``{r0} \ ?ordLess^-1``{r0})" + let ?R = "?ordLess Int (?ordLess\``{r0} \ ?ordLess\``{r0})" {assume Case1: "Well_order r0" hence "wf ?R" using wf_ofilterIncl[of r0] @@ -934,7 +934,7 @@ hence "b1 \ Field r \ b2 \ Field r" unfolding Field_def by auto hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto - hence "(a,c): r" using 2 TRANS unfolding trans_def by blast + hence "(a,c) \ r" using 2 TRANS unfolding trans_def by blast thus "(a',c') \ dir_image r f" unfolding dir_image_def using 1 by auto qed diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/BNF_Wellorder_Embedding.thy --- a/src/HOL/BNF_Wellorder_Embedding.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/BNF_Wellorder_Embedding.thy Thu Feb 15 13:04:36 2018 +0100 @@ -317,10 +317,10 @@ using 2 by auto with Field_def[of r] obtain b where 3: "b \ Field r" and 4: "b' = f b" by auto - have "(b,a): r" + have "(b,a) \ r" proof- {assume "(a,b) \ r" - with ** 4 have "(f a, b'): r'" + with ** 4 have "(f a, b') \ r'" by (auto simp add: compat_def) with ***** Antisym' have "f a = b'" by(auto simp add: under_def antisym_def) @@ -799,7 +799,7 @@ shows "a \ Field r \ f a = g a" proof(rule wo_rel.well_order_induct[of r], auto simp add: WELL wo_rel_def) fix a - assume IH: "\b. b \ a \ (b,a): r \ b \ Field r \ f b = g b" and + assume IH: "\b. b \ a \ (b,a) \ r \ b \ Field r \ f b = g b" and *: "a \ Field r" hence "\b \ underS r a. f b = g b" unfolding underS_def by (auto simp add: Field_def) @@ -1071,7 +1071,7 @@ fix a b assume **: "a \ Field r" "b \ Field r" and ***: "(f a, f b) \ r'" {assume "(b,a) \ r \ b = a" - hence "(b,a): r"using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast + hence "(b,a) \ r"using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast hence "(f b, f a) \ r'" using * unfolding compat_def by auto hence "f a = f b" using Well *** wo_rel.ANTISYM[of r'] antisym_def[of r'] by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/AxCompl.thy --- a/src/HOL/Bali/AxCompl.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/AxCompl.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1376,7 +1376,7 @@ apply (simp add: MGF_asm ax_derivs_asm) apply (rule MGF_asm) apply (rule ballI) -apply (case_tac "mgf_call pn : A") +apply (case_tac "mgf_call pn \ A") apply (fast intro: ax_derivs_asm) apply (rule MGF_nested_Methd) apply (rule ballI) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/AxSem.thy --- a/src/HOL/Bali/AxSem.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/AxSem.thy Thu Feb 15 13:04:36 2018 +0100 @@ -687,7 +687,7 @@ done lemma weaken: - "G,(A::'a triple set)|\(ts'::'a triple set) \ !ts. ts \ ts' \ G,A|\ts" + "G,(A::'a triple set)|\(ts'::'a triple set) \ \ts. ts \ ts' \ G,A|\ts" apply (erule ax_derivs.induct) (*42 subgoals*) apply (tactic "ALLGOALS (strip_tac @{context})") diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/AxSound.thy --- a/src/HOL/Bali/AxSound.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/AxSound.thy Thu Feb 15 13:04:36 2018 +0100 @@ -300,7 +300,7 @@ cases. Auto will then solve premise 6 and 7. *) -lemma all_empty: "(!x. P) = P" +lemma all_empty: "(\x. P) = P" by simp corollary evaln_type_sound: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/Basis.thy Thu Feb 15 13:04:36 2018 +0100 @@ -121,7 +121,7 @@ obtains x s where "s' = (x,s)" and "x = x'" using assms by (cases s') auto -lemma fst_in_set_lemma: "(x, y) : set l \ x : fst ` set l" +lemma fst_in_set_lemma: "(x, y) \ set l \ x \ fst ` set l" by (induct l) auto @@ -246,7 +246,7 @@ lemma unique_map_inj: "unique l \ inj f \ unique (map (\(k,x). (f k, g k x)) l)" by (induct l) (auto dest: fst_in_set_lemma simp add: inj_eq) -lemma map_of_SomeI: "unique l \ (k, x) : set l \ map_of l k = Some x" +lemma map_of_SomeI: "unique l \ (k, x) \ set l \ map_of l k = Some x" by (induct l) auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/Decl.thy --- a/src/HOL/Bali/Decl.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/Decl.thy Thu Feb 15 13:04:36 2018 +0100 @@ -461,11 +461,11 @@ abbreviation subclseq_syntax :: "prog => [qtname, qtname] => bool" ("_\_\\<^sub>C _" [71,71,71] 70) - where "G\C \\<^sub>C D == (C,D) \(subcls1 G)^*" (* cf. 8.1.3 *) + where "G\C \\<^sub>C D == (C,D) \(subcls1 G)\<^sup>*" (* cf. 8.1.3 *) abbreviation subcls_syntax :: "prog => [qtname, qtname] => bool" ("_\_\\<^sub>C _" [71,71,71] 70) - where "G\C \\<^sub>C D == (C,D) \(subcls1 G)^+" + where "G\C \\<^sub>C D == (C,D) \(subcls1 G)\<^sup>+" notation (ASCII) subcls1_syntax ("_|-_<:C1_" [71,71,71] 70) and @@ -520,11 +520,11 @@ definition ws_idecl :: "prog \ qtname \ qtname list \ bool" - where "ws_idecl G I si = (\J\set si. is_iface G J \ (J,I)\(subint1 G)^+)" + where "ws_idecl G I si = (\J\set si. is_iface G J \ (J,I)\(subint1 G)\<^sup>+)" definition ws_cdecl :: "prog \ qtname \ qtname \ bool" - where "ws_cdecl G C sc = (C\Object \ is_class G sc \ (sc,C)\(subcls1 G)^+)" + where "ws_cdecl G C sc = (C\Object \ is_class G sc \ (sc,C)\(subcls1 G)\<^sup>+)" definition ws_prog :: "prog \ bool" where @@ -534,9 +534,9 @@ lemma ws_progI: "\\(I,i)\set (ifaces G). \J\set (isuperIfs i). is_iface G J \ - (J,I) \ (subint1 G)^+; + (J,I) \ (subint1 G)\<^sup>+; \(C,c)\set (classes G). C\Object \ is_class G (super c) \ - ((super c),C) \ (subcls1 G)^+ + ((super c),C) \ (subcls1 G)\<^sup>+ \ \ ws_prog G" apply (unfold ws_prog_def ws_idecl_def ws_cdecl_def) apply (erule_tac conjI) @@ -545,7 +545,7 @@ lemma ws_prog_ideclD: "\iface G I = Some i; J\set (isuperIfs i); ws_prog G\ \ - is_iface G J \ (J,I)\(subint1 G)^+" + is_iface G J \ (J,I)\(subint1 G)\<^sup>+" apply (unfold ws_prog_def ws_idecl_def) apply clarify apply (drule_tac map_of_SomeD) @@ -554,7 +554,7 @@ lemma ws_prog_cdeclD: "\class G C = Some c; C\Object; ws_prog G\ \ - is_class G (super c) \ (super c,C)\(subcls1 G)^+" + is_class G (super c) \ (super c,C)\(subcls1 G)\<^sup>+" apply (unfold ws_prog_def ws_cdecl_def) apply clarify apply (drule_tac map_of_SomeD) @@ -590,12 +590,12 @@ done lemma subint1_irrefl_lemma1: - "ws_prog G \ (subint1 G)^-1 \ (subint1 G)^+ = {}" + "ws_prog G \ (subint1 G)\ \ (subint1 G)\<^sup>+ = {}" apply (force dest: subint1D ws_prog_ideclD conjunct2) done lemma subcls1_irrefl_lemma1: - "ws_prog G \ (subcls1 G)^-1 \ (subcls1 G)^+ = {}" + "ws_prog G \ (subcls1 G)\ \ (subcls1 G)\<^sup>+ = {}" apply (force dest: subcls1D ws_prog_cdeclD conjunct2) done @@ -778,7 +778,7 @@ else undefined)" by auto termination -by (relation "inv_image (same_fst ws_prog (\G. (subint1 G)^-1)) (%(x,y,z). (x,y))") +by (relation "inv_image (same_fst ws_prog (\G. (subint1 G)\)) (%(x,y,z). (x,y))") (auto simp: wf_subint1 subint1I wf_same_fst) lemma iface_rec: @@ -803,7 +803,7 @@ by auto termination -by (relation "inv_image (same_fst ws_prog (\G. (subcls1 G)^-1)) (%(x,y,z,w). (x,y))") +by (relation "inv_image (same_fst ws_prog (\G. (subcls1 G)\)) (%(x,y,z,w). (x,y))") (auto simp: wf_subcls1 subcls1I wf_same_fst) lemma class_rec: "\class G C = Some c; ws_prog G\ \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/DeclConcepts.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1565,7 +1565,7 @@ lemma imethds_declI: "\m \ imethds G I sig; ws_prog G; is_iface G I\ \ (\i. iface G (decliface m) = Some i \ table_of (imethods i) sig = Some (mthd m)) \ - (I,decliface m) \ (subint1 G)^* \ m \ imethds G (decliface m) sig" + (I,decliface m) \ (subint1 G)\<^sup>* \ m \ imethds G (decliface m) sig" apply (erule rev_mp) apply (rule ws_subint1_induct, assumption, assumption) apply (subst imethds_rec, erule conjunct1, assumption) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/Example.thy --- a/src/HOL/Bali/Example.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/Example.thy Thu Feb 15 13:04:36 2018 +0100 @@ -287,30 +287,30 @@ subsubsection "well-structuredness" -lemma not_Object_subcls_any [elim!]: "(Object, C) \ (subcls1 tprg)^+ \ R" +lemma not_Object_subcls_any [elim!]: "(Object, C) \ (subcls1 tprg)\<^sup>+ \ R" apply (auto dest!: tranclD subcls1D) done lemma not_Throwable_subcls_SXcpt [elim!]: - "(SXcpt Throwable, SXcpt xn) \ (subcls1 tprg)^+ \ R" + "(SXcpt Throwable, SXcpt xn) \ (subcls1 tprg)\<^sup>+ \ R" apply (auto dest!: tranclD subcls1D) apply (simp add: Object_def SXcpt_def) done lemma not_SXcpt_n_subcls_SXcpt_n [elim!]: - "(SXcpt xn, SXcpt xn) \ (subcls1 tprg)^+ \ R" + "(SXcpt xn, SXcpt xn) \ (subcls1 tprg)\<^sup>+ \ R" apply (auto dest!: tranclD subcls1D) apply (drule rtranclD) apply auto done -lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \ (subcls1 tprg)^+ \ R" +lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \ (subcls1 tprg)\<^sup>+ \ R" apply (auto dest!: tranclD subcls1D simp add: BaseCl_def) done lemma not_TName_n_subcls_TName_n [rule_format (no_asm), elim!]: "(\pid=java_lang,tid=TName tn\, \pid=java_lang,tid=TName tn\) - \ (subcls1 tprg)^+ \ R" + \ (subcls1 tprg)\<^sup>+ \ R" apply (rule_tac n1 = "tn" in surj_tnam' [THEN exE]) apply (erule ssubst) apply (rule tnam'.induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/Table.thy --- a/src/HOL/Bali/Table.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/Table.thy Thu Feb 15 13:04:36 2018 +0100 @@ -364,10 +364,10 @@ primrec atleast_free :: "('a \ 'b) => nat => bool" where "atleast_free m 0 = True" -| atleast_free_Suc: "atleast_free m (Suc n) = (\a. m a = None & (!b. atleast_free (m(a|->b)) n))" +| atleast_free_Suc: "atleast_free m (Suc n) = (\a. m a = None \ (\b. atleast_free (m(a\b)) n))" lemma atleast_free_weaken [rule_format (no_asm)]: - "!m. atleast_free m (Suc n) \ atleast_free m n" + "\m. atleast_free m (Suc n) \ atleast_free m n" apply (induct_tac "n") apply (simp (no_asm)) apply clarify @@ -377,13 +377,13 @@ done lemma atleast_free_SucI: -"[| h a = None; !obj. atleast_free (h(a|->obj)) n |] ==> atleast_free h (Suc n)" +"[| h a = None; \obj. atleast_free (h(a|->obj)) n |] ==> atleast_free h (Suc n)" by force declare fun_upd_apply [simp del] lemma atleast_free_SucD_lemma [rule_format (no_asm)]: -" !m a. m a = None --> (!c. atleast_free (m(a|->c)) n) --> - (!b d. a ~= b --> atleast_free (m(b|->d)) n)" +"\m a. m a = None \ (\c. atleast_free (m(a\c)) n) \ + (\b d. a \ b \ atleast_free (m(b\d)) n)" apply (induct_tac "n") apply auto apply (rule_tac x = "a" in exI) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/TypeRel.thy --- a/src/HOL/Bali/TypeRel.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/TypeRel.thy Thu Feb 15 13:04:36 2018 +0100 @@ -43,7 +43,7 @@ abbreviation subint_syntax :: "prog => [qtname, qtname] => bool" ("_\_\I _" [71,71,71] 70) - where "G\I \I J == (I,J) \(subint1 G)^*" \ \cf. 9.1.3\ + where "G\I \I J == (I,J) \(subint1 G)\<^sup>*" \ \cf. 9.1.3\ abbreviation implmt1_syntax :: "prog => [qtname, qtname] => bool" ("_\_\1_" [71,71,71] 70) @@ -116,7 +116,7 @@ apply (auto intro: rtrancl_into_trancl3) done -lemma subcls_is_class: "(C,D) \ (subcls1 G)^+ \ is_class G C" +lemma subcls_is_class: "(C,D) \ (subcls1 G)\<^sup>+ \ is_class G C" apply (erule trancl_trans_induct) apply (auto dest!: subcls1D) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Bali/WellForm.thy --- a/src/HOL/Bali/WellForm.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Bali/WellForm.thy Thu Feb 15 13:04:36 2018 +0100 @@ -268,7 +268,7 @@ lemma wf_idecl_supD: "\wf_idecl G (I,i); J \ set (isuperIfs i)\ - \ is_acc_iface G (pid I) J \ (J, I) \ (subint1 G)^+" + \ is_acc_iface G (pid I) J \ (J, I) \ (subint1 G)\<^sup>+" apply (unfold wf_idecl_def ws_idecl_def) apply auto done @@ -451,7 +451,7 @@ lemma wf_cdecl_supD: "\wf_cdecl G (C,c); C \ Object\ \ - is_acc_class G (pid C) (super c) \ (super c,C) \ (subcls1 G)^+ \ + is_acc_class G (pid C) (super c) \ (super c,C) \ (subcls1 G)\<^sup>+ \ (table_of (map (\ (s,m). (s,C,m)) (methods c)) entails (\ new. \ old sig. (G,sig\new overrides\<^sub>S old diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Cardinals/Cardinal_Order_Relation.thy --- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Thu Feb 15 13:04:36 2018 +0100 @@ -71,18 +71,18 @@ subsection \Cardinal of a set\ -lemma card_of_inj_rel: assumes INJ: "!! x y y'. \(x,y) : R; (x,y') : R\ \ y = y'" -shows "|{y. EX x. (x,y) : R}| <=o |{x. EX y. (x,y) : R}|" +lemma card_of_inj_rel: assumes INJ: "\x y y'. \(x,y) \ R; (x,y') \ R\ \ y = y'" +shows "|{y. \x. (x,y) \ R}| <=o |{x. \y. (x,y) \ R}|" proof- - let ?Y = "{y. EX x. (x,y) : R}" let ?X = "{x. EX y. (x,y) : R}" - let ?f = "% y. SOME x. (x,y) : R" + let ?Y = "{y. \x. (x,y) \ R}" let ?X = "{x. \y. (x,y) \ R}" + let ?f = "\y. SOME x. (x,y) \ R" have "?f ` ?Y <= ?X" by (auto dest: someI) moreover have "inj_on ?f ?Y" unfolding inj_on_def proof(auto) fix y1 x1 y2 x2 assume *: "(x1, y1) \ R" "(x2, y2) \ R" and **: "?f y1 = ?f y2" - hence "(?f y1,y1) : R" using someI[of "% x. (x,y1) : R"] by auto - moreover have "(?f y2,y2) : R" using * someI[of "% x. (x,y2) : R"] by auto + hence "(?f y1,y1) \ R" using someI[of "\x. (x,y1) \ R"] by auto + moreover have "(?f y2,y2) \ R" using * someI[of "\x. (x,y2) \ R"] by auto ultimately show "y1 = y2" using ** INJ by auto qed ultimately show "|?Y| <=o |?X|" using card_of_ordLeq by blast @@ -1629,7 +1629,7 @@ shows "\ i \ Field r. \ b \ B. P i b" proof- let ?As = "\i. {b \ B. P i b}" - have "EX i : Field r. B \ ?As i" + have "\i \ Field r. B \ ?As i" apply(rule regularCard_UNION) using assms unfolding relChain_def by auto thus ?thesis by auto qed diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Cardinals/Wellorder_Constructions.thy --- a/src/HOL/Cardinals/Wellorder_Constructions.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Cardinals/Wellorder_Constructions.thy Thu Feb 15 13:04:36 2018 +0100 @@ -37,7 +37,7 @@ "Restr (Restr r A) B = Restr r (A Int B)" by blast -lemma Restr_iff: "(a,b) : Restr r A = (a : A \ b : A \ (a,b) : r)" +lemma Restr_iff: "(a,b) \ Restr r A = (a \ A \ b \ A \ (a,b) \ r)" by (auto simp add: Field_def) lemma Restr_subset1: "Restr r A \ r" @@ -281,7 +281,7 @@ definition isOmax :: "'a rel set \ 'a rel \ bool" where -"isOmax R r == r \ R \ (ALL r' : R. r' \o r)" +"isOmax R r \ r \ R \ (\r' \ R. r' \o r)" definition omax :: "'a rel set \ 'a rel" where @@ -405,8 +405,8 @@ shows "omax P \o omax R" proof- let ?mp = "omax P" let ?mr = "omax R" - {fix p assume "p : P" - then obtain r where r: "r : R" and "p \o r" + {fix p assume "p \ P" + then obtain r where r: "r \ R" and "p \o r" using LEQ by blast moreover have "r <=o ?mr" using r R Well_R omax_maxim by blast @@ -424,8 +424,8 @@ shows "omax P P" + then obtain r where r: "r \ R" and "p {} \ (INF i:I. inf x (f i)) = inf x (INF i:I. f i)" +lemma INF_inf_const1: "I \ {} \ (\i\I. inf x (f i)) = inf x (\i\I. f i)" by (intro antisym INF_greatest inf_mono order_refl INF_lower) (auto intro: INF_lower2 le_infI2 intro!: INF_mono) -lemma INF_inf_const2: "I \ {} \ (INF i:I. inf (f i) x) = inf (INF i:I. f i) x" +lemma INF_inf_const2: "I \ {} \ (\i\I. inf (f i) x) = inf (\i\I. f i) x" using INF_inf_const1[of I x f] by (simp add: inf_commute) lemma INF_constant: "(\y\A. c) = (if A = {} then \ else c)" @@ -514,10 +514,10 @@ lemma mono_Sup: "(\x\A. f x) \ f (\A)" using \mono f\ by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD) -lemma mono_INF: "f (INF i : I. A i) \ (INF x : I. f (A x))" +lemma mono_INF: "f (\i\I. A i) \ (\x\I. f (A x))" by (intro complete_lattice_class.INF_greatest monoD[OF \mono f\] INF_lower) -lemma mono_SUP: "(SUP x : I. f (A x)) \ f (SUP i : I. A i)" +lemma mono_SUP: "(\x\I. f (A x)) \ f (\i\I. A i)" by (intro complete_lattice_class.SUP_least monoD[OF \mono f\] SUP_upper) end @@ -1250,7 +1250,7 @@ lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" by (fact Sup_inf_eq_bot_iff) -lemma SUP_UNION: "(SUP x:(UN y:A. g y). f x) = (SUP y:A. SUP x:g y. f x :: _ :: complete_lattice)" +lemma SUP_UNION: "(\x\(\y\A. g y). f x) = (\y\A. \x\g y. f x :: _ :: complete_lattice)" by (rule order_antisym) (blast intro: SUP_least SUP_upper2)+ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Computational_Algebra/Primes.thy --- a/src/HOL/Computational_Algebra/Primes.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Computational_Algebra/Primes.thy Thu Feb 15 13:04:36 2018 +0100 @@ -728,7 +728,7 @@ define g where "g = (\x. if x \ S then f x else 0)" define A where "A = Abs_multiset g" have "{x. g x > 0} \ S" by (auto simp: g_def) - from finite_subset[OF this assms(1)] have [simp]: "g : multiset" + from finite_subset[OF this assms(1)] have [simp]: "g \ multiset" by (simp add: multiset_def) from assms have count_A: "count A x = g x" for x unfolding A_def by simp diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Thu Feb 15 13:04:36 2018 +0100 @@ -304,10 +304,10 @@ lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ SUPREMUM A f" by (auto intro: cSUP_upper order_trans) -lemma cSUP_const [simp]: "A \ {} \ (SUP x:A. c) = c" +lemma cSUP_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cSUP_least) (auto intro: cSUP_upper) -lemma cINF_const [simp]: "A \ {} \ (INF x:A. c) = c" +lemma cINF_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cINF_greatest) (auto intro: cINF_lower) lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ INFIMUM A f \ (\x\A. u \ f x)" @@ -316,10 +316,10 @@ lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ SUPREMUM A f \ u \ (\x\A. f x \ u)" by (metis cSUP_least cSUP_upper order_trans) -lemma less_cINF_D: "bdd_below (f`A) \ y < (INF i:A. f i) \ i \ A \ y < f i" +lemma less_cINF_D: "bdd_below (f`A) \ y < (\i\A. f i) \ i \ A \ y < f i" by (metis cINF_lower less_le_trans) -lemma cSUP_lessD: "bdd_above (f`A) \ (SUP i:A. f i) < y \ i \ A \ f i < y" +lemma cSUP_lessD: "bdd_above (f`A) \ (\i\A. f i) < y \ i \ A \ f i < y" by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ INFIMUM (insert a A) f = inf (f a) (INFIMUM A f)" @@ -358,11 +358,11 @@ lemma cSUP_union: "A \ {} \ bdd_above (f`A) \ B \ {} \ bdd_above (f`B) \ SUPREMUM (A \ B) f = sup (SUPREMUM A f) (SUPREMUM B f)" using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un [symmetric]) -lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ inf (INFIMUM A f) (INFIMUM A g) = (INF a:A. inf (f a) (g a))" +lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ inf (INFIMUM A f) (INFIMUM A g) = (\a\A. inf (f a) (g a))" by (intro antisym le_infI cINF_greatest cINF_lower2) (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) -lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ sup (SUPREMUM A f) (SUPREMUM A g) = (SUP a:A. sup (f a) (g a))" +lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ sup (SUPREMUM A f) (SUPREMUM A g) = (\a\A. sup (f a) (g a))" by (intro antisym le_supI cSUP_least cSUP_upper2) (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) @@ -403,10 +403,10 @@ lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) -lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (INF i:A. f i) < a \ (\x\A. f x < a)" +lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (\i\A. f i) < a \ (\x\A. f x < a)" using cInf_less_iff[of "f`A"] by auto -lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (SUP i:A. f i) \ (\x\A. a < f x)" +lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (\i\A. f i) \ (\x\A. a < f x)" using less_cSup_iff[of "f`A"] by auto lemma less_cSupE: @@ -642,10 +642,10 @@ lemma cSUP_eq_cINF_D: fixes f :: "_ \ 'b::conditionally_complete_lattice" - assumes eq: "(SUP x:A. f x) = (INF x:A. f x)" + assumes eq: "(\x\A. f x) = (\x\A. f x)" and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" and a: "a \ A" - shows "f a = (INF x:A. f x)" + shows "f a = (\x\A. f x)" apply (rule antisym) using a bdd apply (auto simp: cINF_lower) @@ -656,17 +656,17 @@ fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_above (\x\A. f ` B x)" - shows "(SUP z : \x\A. B x. f z) = (SUP x:A. SUP z:B x. f z)" + shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_above (f ` B x)" using bdd_UN by (meson UN_upper bdd_above_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_above_def) - then have bdd2: "bdd_above ((\x. SUP z:B x. f z) ` A)" + then have bdd2: "bdd_above ((\x. \z\B x. f z) ` A)" unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) - have "(SUP z:\x\A. B x. f z) \ (SUP x:A. SUP z:B x. f z)" + have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) - moreover have "(SUP x:A. SUP z:B x. f z) \ (SUP z:\x\A. B x. f z)" + moreover have "(\x\A. \z\B x. f z) \ (\ z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) @@ -676,17 +676,17 @@ fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_below (\x\A. f ` B x)" - shows "(INF z : \x\A. B x. f z) = (INF x:A. INF z:B x. f z)" + shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_below (f ` B x)" using bdd_UN by (meson UN_upper bdd_below_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_below_def) - then have bdd2: "bdd_below ((\x. INF z:B x. f z) ` A)" + then have bdd2: "bdd_below ((\x. \z\B x. f z) ` A)" unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) - have "(INF z:\x\A. B x. f z) \ (INF x:A. INF z:B x. f z)" + have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) - moreover have "(INF x:A. INF z:B x. f z) \ (INF z:\x\A. B x. f z)" + moreover have "(\x\A. \z\B x. f z) \ (\z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) ultimately show ?thesis by (rule order_antisym) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Data_Structures/AA_Map.thy --- a/src/HOL/Data_Structures/AA_Map.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Data_Structures/AA_Map.thy Thu Feb 15 13:04:36 2018 +0100 @@ -64,7 +64,7 @@ qed simp lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \ - (EX l x r. update a b t = Node (lvl t + 1) l x r \ lvl l = lvl r)" + (\l x r. update a b t = Node (lvl t + 1) l x r \ lvl l = lvl r)" apply(cases t) apply(auto simp add: skew_case split_case split: if_splits) apply(auto split: tree.splits if_splits) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Data_Structures/AA_Set.thy --- a/src/HOL/Data_Structures/AA_Set.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Data_Structures/AA_Set.thy Thu Feb 15 13:04:36 2018 +0100 @@ -193,7 +193,7 @@ lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \ - (EX l x r. insert a t = Node (lvl t + 1) l x r \ lvl l = lvl r)" + (\l x r. insert a t = Node (lvl t + 1) l x r \ lvl l = lvl r)" apply(cases t) apply(auto simp add: skew_case split_case split: if_splits) apply(auto split: tree.splits if_splits) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Data_Structures/Brother12_Set.thy --- a/src/HOL/Data_Structures/Brother12_Set.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Data_Structures/Brother12_Set.thy Thu Feb 15 13:04:36 2018 +0100 @@ -456,7 +456,7 @@ { case 2 with Suc.IH(1) show ?case by auto } qed auto -lemma tree_type: "t \ T (h+1) \ tree t : B (h+1) \ B h" +lemma tree_type: "t \ T (h+1) \ tree t \ B (h+1) \ B h" by(auto) lemma delete_type: "t \ B h \ delete x t \ B h \ B(h-1)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Thu Feb 15 13:04:36 2018 +0100 @@ -2755,17 +2755,17 @@ shows "\ b\ set (\ p). isint b bs" using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) -lemma cpmi_eq: "0 < D \ (EX z::int. ALL x. x < z --> (P x = P1 x)) -==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) -==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) -==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" +lemma cpmi_eq: "0 < D \ (\z::int. \x. x < z \ (P x = P1 x)) +\ \x. \(\(j::int) \ {1..D}. \(b::int) \ B. P(b+j)) \ P (x) \ P (x - D) +\ (\(x::int). \(k::int). ((P1 x)= (P1 (x-k*D)))) +\ (\(x::int). P(x)) = ((\(j::int) \ {1..D} . (P1(j))) | (\(j::int) \ {1..D}. \(b::int) \ B. P (b+j)))" apply(rule iffI) prefer 2 apply(drule minusinfinity) apply assumption+ apply(fastforce) apply clarsimp -apply(subgoal_tac "!!k. 0<=k \ !x. P x \ P (x - k*D)") +apply(subgoal_tac "\k. 0<=k \ \x. P x \ P (x - k*D)") apply(frule_tac x = x and z=z in decr_lemma) apply(subgoal_tac "P1(x - (\x - z\ + 1) * D)") prefer 2 @@ -5710,4 +5710,3 @@ by mir end - diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Decision_Procs/approximation.ML --- a/src/HOL/Decision_Procs/approximation.ML Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Decision_Procs/approximation.ML Thu Feb 15 13:04:36 2018 +0100 @@ -105,9 +105,9 @@ fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t | calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t - | calculated_subterms (@{term "(<=) :: real \ real \ bool"} $ t1 $ t2) = [t1, t2] + | calculated_subterms (@{term "(\) :: real \ real \ bool"} $ t1 $ t2) = [t1, t2] | calculated_subterms (@{term "(<) :: real \ real \ bool"} $ t1 $ t2) = [t1, t2] - | calculated_subterms (@{term "(:) :: real \ real set \ bool"} $ t1 $ + | calculated_subterms (@{term "(\) :: real \ real set \ bool"} $ t1 $ (@{term "atLeastAtMost :: real \ real \ real set"} $ t2 $ t3)) = [t1, t2, t3] | calculated_subterms t = raise TERM ("calculated_subterms", [t]) @@ -246,7 +246,7 @@ |> approximate ctxt |> HOLogic.dest_list |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term) - |> map (fn (elem, s) => @{term "(:) :: real \ real set \ bool"} $ elem $ mk_result prec (dest_ivl s)) + |> map (fn (elem, s) => @{term "(\) :: real \ real set \ bool"} $ elem $ mk_result prec (dest_ivl s)) |> foldr1 HOLogic.mk_conj)) fun approx_arith prec ctxt t = realify t diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Enum.thy --- a/src/HOL/Enum.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Enum.thy Thu Feb 15 13:04:36 2018 +0100 @@ -61,7 +61,7 @@ by (simp add: enum_UNIV) lemma vimage_code [code]: - "f -` B = set (filter (%x. f x : B) enum_class.enum)" + "f -` B = set (filter (\x. f x \ B) enum_class.enum)" unfolding vimage_def Collect_code .. definition card_UNIV :: "'a itself \ nat" @@ -208,8 +208,8 @@ shows "Wellfounded.acc r \ (\n. bacc r n)" proof fix x - assume "x : Wellfounded.acc r" - then have "\ n. x : bacc r n" + assume "x \ Wellfounded.acc r" + then have "\n. x \ bacc r n" proof (induct x arbitrary: rule: acc.induct) case (accI x) then have "\y. \ n. (y, x) \ r \ y \ bacc r n" by simp @@ -217,16 +217,16 @@ obtain n where "\y. (y, x) \ r \ y \ bacc r n" proof fix y assume y: "(y, x) \ r" - with n have "y : bacc r (n y)" by auto + with n have "y \ bacc r (n y)" by auto moreover have "n y <= Max ((\(y, x). n y) ` r)" using y \finite r\ by (auto intro!: Max_ge) note bacc_mono[OF this, of r] - ultimately show "y : bacc r (Max ((\(y, x). n y) ` r))" by auto + ultimately show "y \ bacc r (Max ((\(y, x). n y) ` r))" by auto qed then show ?case by (auto simp add: Let_def intro!: exI[of _ "Suc n"]) qed - then show "x : (UN n. bacc r n)" by auto + then show "x \ (\n. bacc r n)" by auto qed lemma acc_bacc_eq: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Filter.thy --- a/src/HOL/Filter.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Filter.thy Thu Feb 15 13:04:36 2018 +0100 @@ -467,7 +467,7 @@ by (simp add: eventually_F) qed -lemma eventually_INF: "eventually P (INF b:B. F b) \ (\X\B. finite X \ eventually P (INF b:X. F b))" +lemma eventually_INF: "eventually P (\b\B. F b) \ (\X\B. finite X \ eventually P (\b\X. F b))" unfolding eventually_Inf [of P "F`B"] by (metis finite_imageI image_mono finite_subset_image) @@ -479,7 +479,7 @@ lemma INF_filter_not_bot: fixes F :: "'i \ 'a filter" - shows "(\X. X \ B \ finite X \ (INF b:X. F b) \ bot) \ (INF b:B. F b) \ bot" + shows "(\X. X \ B \ finite X \ (\b\X. F b) \ bot) \ (\b\B. F b) \ bot" unfolding trivial_limit_def eventually_INF [of _ _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp @@ -507,10 +507,10 @@ lemma eventually_INF_base: "B \ {} \ (\a b. a \ B \ b \ B \ \x\B. F x \ inf (F a) (F b)) \ - eventually P (INF b:B. F b) \ (\b\B. eventually P (F b))" + eventually P (\b\B. F b) \ (\b\B. eventually P (F b))" by (subst eventually_Inf_base) auto -lemma eventually_INF1: "i \ I \ eventually P (F i) \ eventually P (INF i:I. F i)" +lemma eventually_INF1: "i \ I \ eventually P (F i) \ eventually P (\i\I. F i)" using filter_leD[OF INF_lower] . lemma eventually_INF_mono: @@ -569,15 +569,15 @@ lemma filtermap_inf: "filtermap f (inf F1 F2) \ inf (filtermap f F1) (filtermap f F2)" by (auto simp: le_filter_def eventually_filtermap eventually_inf) -lemma filtermap_INF: "filtermap f (INF b:B. F b) \ (INF b:B. filtermap f (F b))" +lemma filtermap_INF: "filtermap f (\b\B. F b) \ (\b\B. filtermap f (F b))" proof - { fix X :: "'c set" assume "finite X" - then have "filtermap f (INFIMUM X F) \ (INF b:X. filtermap f (F b))" + then have "filtermap f (INFIMUM X F) \ (\b\X. filtermap f (F b))" proof induct case (insert x X) - have "filtermap f (INF a:insert x X. F a) \ inf (filtermap f (F x)) (filtermap f (INF a:X. F a))" + have "filtermap f (\a\insert x X. F a) \ inf (filtermap f (F x)) (filtermap f (\a\X. F a))" by (rule order_trans[OF _ filtermap_inf]) simp - also have "\ \ inf (filtermap f (F x)) (INF a:X. filtermap f (F a))" + also have "\ \ inf (filtermap f (F x)) (\a\X. filtermap f (F a))" by (intro inf_mono insert order_refl) finally show ?case by simp @@ -660,17 +660,17 @@ by (auto simp: filter_eq_iff eventually_filtercomap eventually_sup) qed -lemma filtercomap_INF: "filtercomap f (INF b:B. F b) = (INF b:B. filtercomap f (F b))" +lemma filtercomap_INF: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" proof - - have *: "filtercomap f (INF b:B. F b) = (INF b:B. filtercomap f (F b))" if "finite B" for B + have *: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" if "finite B" for B using that by induction (simp_all add: filtercomap_inf) show ?thesis unfolding filter_eq_iff proof fix P - have "eventually P (INF b:B. filtercomap f (F b)) \ + have "eventually P (\b\B. filtercomap f (F b)) \ (\X. (X \ B \ finite X) \ eventually P (\b\X. filtercomap f (F b)))" by (subst eventually_INF) blast - also have "\ \ (\X. (X \ B \ finite X) \ eventually P (filtercomap f (INF b:X. F b)))" + also have "\ \ (\X. (X \ B \ finite X) \ eventually P (filtercomap f (\b\X. F b)))" by (rule ex_cong) (simp add: *) also have "\ \ eventually P (filtercomap f (INFIMUM B F))" unfolding eventually_filtercomap by (subst eventually_INF) blast @@ -680,7 +680,7 @@ qed lemma filtercomap_SUP_finite: - "finite B \ filtercomap f (SUP b:B. F b) \ (SUP b:B. filtercomap f (F b))" + "finite B \ filtercomap f (\b\B. F b) \ (\b\B. filtercomap f (F b))" by (induction B rule: finite_induct) (auto intro: order_trans[OF _ order_trans[OF _ filtercomap_sup]] filtercomap_mono) @@ -738,10 +738,10 @@ unfolding filter_eq_iff eventually_inf eventually_principal by (auto intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) -lemma SUP_principal[simp]: "(SUP i : I. principal (A i)) = principal (\i\I. A i)" +lemma SUP_principal[simp]: "(\i\I. principal (A i)) = principal (\i\I. A i)" unfolding filter_eq_iff eventually_Sup by (auto simp: eventually_principal) -lemma INF_principal_finite: "finite X \ (INF x:X. principal (f x)) = principal (\x\X. f x)" +lemma INF_principal_finite: "finite X \ (\x\X. principal (f x)) = principal (\x\X. f x)" by (induct X rule: finite_induct) auto lemma filtermap_principal[simp]: "filtermap f (principal A) = principal (f ` A)" @@ -753,9 +753,9 @@ subsubsection \Order filters\ definition at_top :: "('a::order) filter" - where "at_top = (INF k. principal {k ..})" + where "at_top = (\k. principal {k ..})" -lemma at_top_sub: "at_top = (INF k:{c::'a::linorder..}. principal {k ..})" +lemma at_top_sub: "at_top = (\k\{c::'a::linorder..}. principal {k ..})" by (auto intro!: INF_eq max.cobounded1 max.cobounded2 simp: at_top_def) lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" @@ -778,9 +778,9 @@ lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::{no_top, linorder}. \n>N. P n)" proof - - have "eventually P (INF k. principal {k <..}) \ (\N::'a. \n>N. P n)" + have "eventually P (\k. principal {k <..}) \ (\N::'a. \n>N. P n)" by (subst eventually_INF_base) (auto simp: eventually_principal intro: max.cobounded1 max.cobounded2) - also have "(INF k. principal {k::'a <..}) = at_top" + also have "(\k. principal {k::'a <..}) = at_top" unfolding at_top_def by (intro INF_eq) (auto intro: less_imp_le simp: Ici_subset_Ioi_iff gt_ex) finally show ?thesis . @@ -806,9 +806,9 @@ qed definition at_bot :: "('a::order) filter" - where "at_bot = (INF k. principal {.. k})" + where "at_bot = (\k. principal {.. k})" -lemma at_bot_sub: "at_bot = (INF k:{.. c::'a::linorder}. principal {.. k})" +lemma at_bot_sub: "at_bot = (\k\{.. c::'a::linorder}. principal {.. k})" by (auto intro!: INF_eq min.cobounded1 min.cobounded2 simp: at_bot_def) lemma eventually_at_bot_linorder: @@ -826,9 +826,9 @@ lemma eventually_at_bot_dense: "eventually P at_bot \ (\N::'a::{no_bot, linorder}. \n (\N::'a. \nk. principal {..< k}) \ (\N::'a. \nk. principal {..< k::'a}) = at_bot" unfolding at_bot_def by (intro INF_eq) (auto intro: less_imp_le simp: Iic_subset_Iio_iff lt_ex) finally show ?thesis . @@ -943,7 +943,7 @@ definition prod_filter :: "'a filter \ 'b filter \ ('a \ 'b) filter" (infixr "\\<^sub>F" 80) where "prod_filter F G = - (INF (P, Q):{(P, Q). eventually P F \ eventually Q G}. principal {(x, y). P x \ Q y})" + (\(P, Q)\{(P, Q). eventually P F \ eventually Q G}. principal {(x, y). P x \ Q y})" lemma eventually_prod_filter: "eventually P (F \\<^sub>F G) \ (\Pf Pg. eventually Pf F \ eventually Pg G \ (\x y. Pf x \ Pg y \ P (x, y)))" @@ -994,16 +994,16 @@ lemma INF_filter_bot_base: fixes F :: "'a \ 'b filter" assumes *: "\i j. i \ I \ j \ I \ \k\I. F k \ F i \ F j" - shows "(INF i:I. F i) = bot \ (\i\I. F i = bot)" + shows "(\i\I. F i) = bot \ (\i\I. F i = bot)" proof (cases "\i\I. F i = bot") case True - then have "(INF i:I. F i) \ bot" + then have "(\i\I. F i) \ bot" by (auto intro: INF_lower2) with True show ?thesis by (auto simp: bot_unique) next case False - moreover have "(INF i:I. F i) \ bot" + moreover have "(\i\I. F i) \ bot" proof (cases "I = {}") case True then show ?thesis @@ -1098,7 +1098,7 @@ lemma prod_filter_INF: assumes "I \ {}" "J \ {}" - shows "(INF i:I. A i) \\<^sub>F (INF j:J. B j) = (INF i:I. INF j:J. A i \\<^sub>F B j)" + shows "(\i\I. A i) \\<^sub>F (\j\J. B j) = (\i\I. \j\J. A i \\<^sub>F B j)" proof (safe intro!: antisym INF_greatest) from \I \ {}\ obtain i where "i \ I" by auto from \J \ {}\ obtain j where "j \ J" by auto @@ -1146,10 +1146,10 @@ unfolding prod_filter_def by (intro eventually_INF1[of "(P, Q)"]) (auto simp: eventually_principal) -lemma prod_filter_INF1: "I \ {} \ (INF i:I. A i) \\<^sub>F B = (INF i:I. A i \\<^sub>F B)" +lemma prod_filter_INF1: "I \ {} \ (\i\I. A i) \\<^sub>F B = (\i\I. A i \\<^sub>F B)" using prod_filter_INF[of I "{B}" A "\x. x"] by simp -lemma prod_filter_INF2: "J \ {} \ A \\<^sub>F (INF i:J. B i) = (INF i:J. A \\<^sub>F B i)" +lemma prod_filter_INF2: "J \ {} \ A \\<^sub>F (\i\J. B i) = (\i\J. A \\<^sub>F B i)" using prod_filter_INF[of "{A}" J "\x. x" B] by simp subsection \Limits\ @@ -1228,21 +1228,21 @@ unfolding filterlim_def by simp lemma filterlim_INF: - "(LIM x F. f x :> (INF b:B. G b)) \ (\b\B. LIM x F. f x :> G b)" + "(LIM x F. f x :> (\b\B. G b)) \ (\b\B. LIM x F. f x :> G b)" unfolding filterlim_def le_INF_iff .. lemma filterlim_INF_INF: - "(\m. m \ J \ \i\I. filtermap f (F i) \ G m) \ LIM x (INF i:I. F i). f x :> (INF j:J. G j)" + "(\m. m \ J \ \i\I. filtermap f (F i) \ G m) \ LIM x (\i\I. F i). f x :> (\j\J. G j)" unfolding filterlim_def by (rule order_trans[OF filtermap_INF INF_mono]) lemma filterlim_base: "(\m x. m \ J \ i m \ I) \ (\m x. m \ J \ x \ F (i m) \ f x \ G m) \ - LIM x (INF i:I. principal (F i)). f x :> (INF j:J. principal (G j))" + LIM x (\i\I. principal (F i)). f x :> (\j\J. principal (G j))" by (force intro!: filterlim_INF_INF simp: image_subset_iff) lemma filterlim_base_iff: assumes "I \ {}" and chain: "\i j. i \ I \ j \ I \ F i \ F j \ F j \ F i" - shows "(LIM x (INF i:I. principal (F i)). f x :> INF j:J. principal (G j)) \ + shows "(LIM x (\i\I. principal (F i)). f x :> \j\J. principal (G j)) \ (\j\J. \i\I. \x\F i. f x \ G j)" unfolding filterlim_INF filterlim_principal proof (subst eventually_INF_base) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Fun_Def.thy --- a/src/HOL/Fun_Def.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Fun_Def.thy Thu Feb 15 13:04:36 2018 +0100 @@ -198,7 +198,7 @@ subsection \Concrete orders for SCNP termination proofs\ definition "pair_less = less_than <*lex*> less_than" -definition "pair_leq = pair_less^=" +definition "pair_leq = pair_less\<^sup>=" definition "max_strict = max_ext pair_less" definition "max_weak = max_ext pair_leq \ {({}, {})}" definition "min_strict = min_ext pair_less" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/FOCUS/Buffer.thy --- a/src/HOL/HOLCF/FOCUS/Buffer.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/FOCUS/Buffer.thy Thu Feb 15 13:04:36 2018 +0100 @@ -97,7 +97,7 @@ "BufSt = {f. \h\BufSt_P. f = h \}" -lemma set_cong: "!!X. A = B ==> (x:A) = (x:B)" +lemma set_cong: "\X. A = B \ (x\A) = (x\B)" by (erule subst, rule refl) @@ -108,21 +108,21 @@ lemmas BufEq_fix = mono_BufEq_F [THEN BufEq_def [THEN eq_reflection, THEN def_gfp_unfold]] -lemma BufEq_unfold: "(f:BufEq) = (!d. f\(Md d\<>) = <> & - (!x. ? ff:BufEq. f\(Md d\\\x) = d\(ff\x)))" +lemma BufEq_unfold: "(f\BufEq) = (\d. f\(Md d\<>) = <> \ + (\x. \ff\BufEq. f\(Md d\\\x) = d\(ff\x)))" apply (subst BufEq_fix [THEN set_cong]) apply (unfold BufEq_F_def) apply (simp) done -lemma Buf_f_empty: "f:BufEq \ f\<> = <>" +lemma Buf_f_empty: "f\BufEq \ f\<> = <>" by (drule BufEq_unfold [THEN iffD1], auto) -lemma Buf_f_d: "f:BufEq \ f\(Md d\<>) = <>" +lemma Buf_f_d: "f\BufEq \ f\(Md d\<>) = <>" by (drule BufEq_unfold [THEN iffD1], auto) lemma Buf_f_d_req: - "f:BufEq \ \ff. ff:BufEq \ f\(Md d\\\x) = d\ff\x" + "f\BufEq \ \ff. ff\BufEq \ f\(Md d\\\x) = d\ff\x" by (drule BufEq_unfold [THEN iffD1], auto) @@ -134,21 +134,21 @@ lemmas BufAC_Asm_fix = mono_BufAC_Asm_F [THEN BufAC_Asm_def [THEN eq_reflection, THEN def_gfp_unfold]] -lemma BufAC_Asm_unfold: "(s:BufAC_Asm) = (s = <> | (? d x. - s = Md d\x & (x = <> | (ft\x = Def \ & (rt\x):BufAC_Asm))))" +lemma BufAC_Asm_unfold: "(s\BufAC_Asm) = (s = <> \ (\d x. + s = Md d\x \ (x = <> \ (ft\x = Def \ \ (rt\x)\BufAC_Asm))))" apply (subst BufAC_Asm_fix [THEN set_cong]) apply (unfold BufAC_Asm_F_def) apply (simp) done -lemma BufAC_Asm_empty: "<> :BufAC_Asm" +lemma BufAC_Asm_empty: "<> \ BufAC_Asm" by (rule BufAC_Asm_unfold [THEN iffD2], auto) -lemma BufAC_Asm_d: "Md d\<>:BufAC_Asm" +lemma BufAC_Asm_d: "Md d\<> \ BufAC_Asm" by (rule BufAC_Asm_unfold [THEN iffD2], auto) -lemma BufAC_Asm_d_req: "x:BufAC_Asm ==> Md d\\\x:BufAC_Asm" +lemma BufAC_Asm_d_req: "x \ BufAC_Asm \ Md d\\\x \ BufAC_Asm" by (rule BufAC_Asm_unfold [THEN iffD2], auto) -lemma BufAC_Asm_prefix2: "a\b\s:BufAC_Asm ==> s:BufAC_Asm" +lemma BufAC_Asm_prefix2: "a\b\s \ BufAC_Asm ==> s \ BufAC_Asm" by (drule BufAC_Asm_unfold [THEN iffD1], auto) @@ -160,31 +160,31 @@ lemmas BufAC_Cmt_fix = mono_BufAC_Cmt_F [THEN BufAC_Cmt_def [THEN eq_reflection, THEN def_gfp_unfold]] -lemma BufAC_Cmt_unfold: "((s,t):BufAC_Cmt) = (!d x. - (s = <> --> t = <>) & - (s = Md d\<> --> t = <>) & - (s = Md d\\\x --> ft\t = Def d & (x, rt\t):BufAC_Cmt))" +lemma BufAC_Cmt_unfold: "((s,t) \ BufAC_Cmt) = (\d x. + (s = <> \ t = <>) \ + (s = Md d\<> \ t = <>) \ + (s = Md d\\\x \ ft\t = Def d \ (x, rt\t) \ BufAC_Cmt))" apply (subst BufAC_Cmt_fix [THEN set_cong]) apply (unfold BufAC_Cmt_F_def) apply (simp) done -lemma BufAC_Cmt_empty: "f:BufEq ==> (<>, f\<>):BufAC_Cmt" +lemma BufAC_Cmt_empty: "f \ BufEq \ (<>, f\<>) \ BufAC_Cmt" by (rule BufAC_Cmt_unfold [THEN iffD2], auto simp add: Buf_f_empty) -lemma BufAC_Cmt_d: "f:BufEq ==> (a\\, f\(a\\)):BufAC_Cmt" +lemma BufAC_Cmt_d: "f \ BufEq \ (a\\, f\(a\\)) \ BufAC_Cmt" by (rule BufAC_Cmt_unfold [THEN iffD2], auto simp add: Buf_f_d) lemma BufAC_Cmt_d2: - "(Md d\\, f\(Md d\\)):BufAC_Cmt ==> f\(Md d\\) = \" + "(Md d\\, f\(Md d\\)) \ BufAC_Cmt \ f\(Md d\\) = \" by (drule BufAC_Cmt_unfold [THEN iffD1], auto) lemma BufAC_Cmt_d3: -"(Md d\\\x, f\(Md d\\\x)):BufAC_Cmt ==> (x, rt\(f\(Md d\\\x))):BufAC_Cmt" +"(Md d\\\x, f\(Md d\\\x)) \ BufAC_Cmt \ (x, rt\(f\(Md d\\\x))) \ BufAC_Cmt" by (drule BufAC_Cmt_unfold [THEN iffD1], auto) lemma BufAC_Cmt_d32: -"(Md d\\\x, f\(Md d\\\x)):BufAC_Cmt ==> ft\(f\(Md d\\\x)) = Def d" +"(Md d\\\x, f\(Md d\\\x)) \ BufAC_Cmt \ ft\(f\(Md d\\\x)) = Def d" by (drule BufAC_Cmt_unfold [THEN iffD1], auto) (**** BufAC *******************************************************************) @@ -194,13 +194,13 @@ apply (fast intro: BufAC_Cmt_d2 BufAC_Asm_d) done -lemma ex_elim_lemma: "(? ff:B. (!x. f\(a\b\x) = d\ff\x)) = - ((!x. ft\(f\(a\b\x)) = Def d) & (LAM x. rt\(f\(a\b\x))):B)" +lemma ex_elim_lemma: "(\ff\B. (\x. f\(a\b\x) = d\ff\x)) = + ((\x. ft\(f\(a\b\x)) = Def d) \ (LAM x. rt\(f\(a\b\x)))\B)" (* this is an instance (though unification cannot handle this) of lemma "(? ff:B. (!x. f\x = d\ff\x)) = \ \((!x. ft\(f\x) = Def d) & (LAM x. rt\(f\x)):B)"*) apply safe -apply ( rule_tac [2] P="(%x. x:B)" in ssubst) +apply ( rule_tac [2] P="(\x. x\B)" in ssubst) prefer 3 apply ( assumption) apply ( rule_tac [2] cfun_eqI) @@ -236,19 +236,19 @@ lemmas BufSt_P_fix = mono_BufSt_F [THEN BufSt_P_def [THEN eq_reflection, THEN def_gfp_unfold]] -lemma BufSt_P_unfold: "(h:BufSt_P) = (!s. h s\<> = <> & - (!d x. h \ \(Md d\x) = h (Sd d)\x & - (? hh:BufSt_P. h (Sd d)\(\\x) = d\(hh \ \x))))" +lemma BufSt_P_unfold: "(h\BufSt_P) = (\s. h s\<> = <> \ + (\d x. h \ \(Md d\x) = h (Sd d)\x \ + (\hh\BufSt_P. h (Sd d)\(\\x) = d\(hh \ \x))))" apply (subst BufSt_P_fix [THEN set_cong]) apply (unfold BufSt_F_def) apply (simp) done -lemma BufSt_P_empty: "h:BufSt_P ==> h s \ <> = <>" +lemma BufSt_P_empty: "h \ BufSt_P \ h s \ <> = <>" by (drule BufSt_P_unfold [THEN iffD1], auto) -lemma BufSt_P_d: "h:BufSt_P ==> h \ \(Md d\x) = h (Sd d)\x" +lemma BufSt_P_d: "h \ BufSt_P \ h \ \(Md d\x) = h (Sd d)\x" by (drule BufSt_P_unfold [THEN iffD1], auto) -lemma BufSt_P_d_req: "h:BufSt_P ==> \hh\BufSt_P. +lemma BufSt_P_d_req: "h \ BufSt_P ==> \hh\BufSt_P. h (Sd d)\(\ \x) = d\(hh \ \x)" by (drule BufSt_P_unfold [THEN iffD1], auto) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/FOCUS/Buffer_adm.thy --- a/src/HOL/HOLCF/FOCUS/Buffer_adm.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/FOCUS/Buffer_adm.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,29 +10,29 @@ declare enat_0 [simp] -lemma BufAC_Asm_d2: "a\s:BufAC_Asm ==> ? d. a=Md d" +lemma BufAC_Asm_d2: "a\s \ BufAC_Asm \ \d. a=Md d" by (drule BufAC_Asm_unfold [THEN iffD1], auto) lemma BufAC_Asm_d3: - "a\b\s:BufAC_Asm ==> ? d. a=Md d & b=\ & s:BufAC_Asm" + "a\b\s \ BufAC_Asm \ \d. a=Md d \ b=\ \ s \ BufAC_Asm" by (drule BufAC_Asm_unfold [THEN iffD1], auto) lemma BufAC_Asm_F_def3: - "(s:BufAC_Asm_F A) = (s=<> | - (? d. ft\s=Def(Md d)) & (rt\s=<> | ft\(rt\s)=Def \ & rt\(rt\s):A))" + "(s \ BufAC_Asm_F A) = (s=<> \ + (\d. ft\s=Def(Md d)) \ (rt\s=<> \ ft\(rt\s)=Def \ \ rt\(rt\s) \ A))" by (unfold BufAC_Asm_F_def, auto) lemma cont_BufAC_Asm_F: "inf_continuous BufAC_Asm_F" by (auto simp add: inf_continuous_def BufAC_Asm_F_def3) lemma BufAC_Cmt_F_def3: - "((s,t):BufAC_Cmt_F C) = (!d x. - (s = <> --> t = <> ) & - (s = Md d\<> --> t = <> ) & - (s = Md d\\\x --> ft\t = Def d & (x,rt\t):C))" + "((s,t) \ BufAC_Cmt_F C) = (\d x. + (s = <> \ t = <> ) \ + (s = Md d\<> \ t = <> ) \ + (s = Md d\\\x \ ft\t = Def d & (x,rt\t) \ C))" apply (unfold BufAC_Cmt_F_def) -apply (subgoal_tac "!d x. (s = Md d\\\x --> (? y. t = d\y & (x,y):C)) = - (s = Md d\\\x --> ft\t = Def d & (x,rt\t):C)") +apply (subgoal_tac "\d x. (s = Md d\\\x \ (\y. t = d\y \ (x,y) \ C)) = + (s = Md d\\\x \ ft\t = Def d \ (x,rt\t) \ C)") apply (simp) apply (auto intro: surjectiv_scons [symmetric]) done @@ -45,12 +45,12 @@ lemma BufAC_Asm_F_stream_monoP: "stream_monoP BufAC_Asm_F" apply (unfold BufAC_Asm_F_def stream_monoP_def) -apply (rule_tac x="{x. (? d. x = Md d\\\<>)}" in exI) +apply (rule_tac x="{x. (\d. x = Md d\\\<>)}" in exI) apply (rule_tac x="Suc (Suc 0)" in exI) apply (clarsimp) done -lemma adm_BufAC_Asm: "adm (%x. x:BufAC_Asm)" +lemma adm_BufAC_Asm: "adm (\x. x \ BufAC_Asm)" apply (unfold BufAC_Asm_def) apply (rule cont_BufAC_Asm_F [THEN BufAC_Asm_F_stream_monoP [THEN fstream_gfp_admI]]) done @@ -61,7 +61,7 @@ lemma BufAC_Asm_F_stream_antiP: "stream_antiP BufAC_Asm_F" apply (unfold stream_antiP_def BufAC_Asm_F_def) apply (intro strip) -apply (rule_tac x="{x. (? d. x = Md d\\\<>)}" in exI) +apply (rule_tac x="{x. (\d. x = Md d\\\<>)}" in exI) apply (rule_tac x="Suc (Suc 0)" in exI) apply (rule conjI) prefer 2 @@ -71,7 +71,7 @@ apply (force)+ done -lemma adm_non_BufAC_Asm: "adm (%u. u~:BufAC_Asm)" +lemma adm_non_BufAC_Asm: "adm (\u. u \ BufAC_Asm)" apply (unfold BufAC_Asm_def) apply (rule cont_BufAC_Asm_F [THEN BufAC_Asm_F_stream_antiP [THEN fstream_non_gfp_admI]]) done @@ -79,7 +79,7 @@ (**** adm_BufAC ***************************************************************) (*adm_non_BufAC_Asm*) -lemma BufAC_Asm_cong [rule_format]: "!f ff. f:BufEq --> ff:BufEq --> s:BufAC_Asm --> f\s = ff\s" +lemma BufAC_Asm_cong [rule_format]: "\f ff. f \ BufEq \ ff \ BufEq \ s \ BufAC_Asm \ f\s = ff\s" apply (rule fstream_ind2) apply (simp add: adm_non_BufAC_Asm) apply (force dest: Buf_f_empty) @@ -92,7 +92,7 @@ (*adm_non_BufAC_Asm,BufAC_Asm_cong*) lemma BufAC_Cmt_d_req: -"!!X. [|f:BufEq; s:BufAC_Asm; (s, f\s):BufAC_Cmt|] ==> (a\b\s, f\(a\b\s)):BufAC_Cmt" +"\X. [|f \ BufEq; s \ BufAC_Asm; (s, f\s) \ BufAC_Cmt|] ==> (a\b\s, f\(a\b\s)) \ BufAC_Cmt" apply (rule BufAC_Cmt_unfold [THEN iffD2]) apply (intro strip) apply (frule Buf_f_d_req) @@ -116,9 +116,9 @@ done (*adm_BufAC_Asm,BufAC_Asm_antiton,adm_non_BufAC_Asm,BufAC_Asm_cong*) -lemma BufAC_Cmt_2stream_monoP: "f:BufEq ==> ? l. !i x s. s:BufAC_Asm --> x << s --> enat (l i) < #x --> - (x,f\x):(BufAC_Cmt_F ^^ i) top --> - (s,f\s):(BufAC_Cmt_F ^^ i) top" +lemma BufAC_Cmt_2stream_monoP: "f \ BufEq \ \l. \i x s. s \ BufAC_Asm \ x << s \ enat (l i) < #x \ + (x,f\x) \ (BufAC_Cmt_F ^^ i) top \ + (s,f\s) \ (BufAC_Cmt_F ^^ i) top" apply (rule_tac x="%i. 2*i" in exI) apply (rule allI) apply (induct_tac "i") @@ -190,7 +190,7 @@ (*adm_BufAC_Asm,BufAC_Asm_antiton,adm_non_BufAC_Asm,BufAC_Asm_cong, BufAC_Cmt_2stream_monoP*) -lemma adm_BufAC: "f:BufEq ==> adm (%s. s:BufAC_Asm --> (s, f\s):BufAC_Cmt)" +lemma adm_BufAC: "f \ BufEq \ adm (\s. s \ BufAC_Asm \ (s, f\s) \ BufAC_Cmt)" apply (rule flatstream_admI) apply (subst BufAC_Cmt_iterate_all) apply (drule BufAC_Cmt_2stream_monoP) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/FOCUS/FOCUS.thy --- a/src/HOL/HOLCF/FOCUS/FOCUS.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/FOCUS/FOCUS.thy Thu Feb 15 13:04:36 2018 +0100 @@ -8,16 +8,16 @@ imports Fstream begin -lemma ex_eqI [intro!]: "? xx. x = xx" +lemma ex_eqI [intro!]: "\xx. x = xx" by auto -lemma ex2_eqI [intro!]: "? xx yy. x = xx & y = yy" +lemma ex2_eqI [intro!]: "\xx yy. x = xx & y = yy" by auto lemma eq_UU_symf: "(UU = f x) = (f x = UU)" by auto -lemma fstream_exhaust_slen_eq: "(#x ~= 0) = (? a y. x = a~> y)" +lemma fstream_exhaust_slen_eq: "(#x \ 0) = (\a y. x = a~> y)" by (simp add: slen_empty_eq fstream_exhaust_eq) lemmas [simp] = diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/FOCUS/Fstream.thy --- a/src/HOL/HOLCF/FOCUS/Fstream.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/FOCUS/Fstream.thy Thu Feb 15 13:04:36 2018 +0100 @@ -52,7 +52,7 @@ apply (simp) done -lemma fstream_exhaust: "x = UU | (? a y. x = a~> y)" +lemma fstream_exhaust: "x = UU \ (\a y. x = a~> y)" apply (simp add: fscons_def2) apply (cut_tac stream.nchotomy) apply (fast dest: not_Undef_is_Def [THEN iffD1]) @@ -65,20 +65,20 @@ apply fast done -lemma fstream_exhaust_eq: "(x ~= UU) = (? a y. x = a~> y)" +lemma fstream_exhaust_eq: "(x \ UU) = (\a y. x = a~> y)" apply (simp add: fscons_def2 stream_exhaust_eq) apply (fast dest: not_Undef_is_Def [THEN iffD1] elim: DefE) done -lemma fscons_not_empty [simp]: "a~> s ~= <>" +lemma fscons_not_empty [simp]: "a~> s \ <>" by (simp add: fscons_def2) -lemma fscons_inject [simp]: "(a~> s = b~> t) = (a = b & s = t)" +lemma fscons_inject [simp]: "(a~> s = b~> t) = (a = b \ s = t)" by (simp add: fscons_def2) -lemma fstream_prefix: "a~> s << t ==> ? tt. t = a~> tt & s << tt" +lemma fstream_prefix: "a~> s << t ==> \tt. t = a~> tt \ s << tt" apply (cases t) apply (cut_tac fscons_not_empty) apply (fast dest: bottomI) @@ -86,7 +86,7 @@ done lemma fstream_prefix' [simp]: - "x << a~> z = (x = <> | (? y. x = a~> y & y << z))" + "x << a~> z = (x = <> \ (\y. x = a~> y \ y << z))" apply (simp add: fscons_def2 lift.distinct(2) [THEN stream_prefix']) apply (safe) apply (erule_tac [!] contrapos_np) @@ -110,7 +110,7 @@ lemma rt_fscons [simp]: "rt\(m~> s) = s" by (simp add: fscons_def) -lemma ft_eq [simp]: "(ft\s = Def a) = (? t. s = a~> t)" +lemma ft_eq [simp]: "(ft\s = Def a) = (\t. s = a~> t)" apply (unfold fscons_def) apply (simp) apply (safe) @@ -144,13 +144,13 @@ by (simp add: fscons_def) lemma slen_fscons_eq: - "(enat (Suc n) < #x) = (? a y. x = a~> y & enat n < #y)" + "(enat (Suc n) < #x) = (\a y. x = a~> y \ enat n < #y)" apply (simp add: fscons_def2 slen_scons_eq) apply (fast dest: not_Undef_is_Def [THEN iffD1] elim: DefE) done lemma slen_fscons_eq_rev: - "(#x < enat (Suc (Suc n))) = (!a y. x ~= a~> y | #y < enat (Suc n))" + "(#x < enat (Suc (Suc n))) = (\a y. x \ a~> y \ #y < enat (Suc n))" apply (simp add: fscons_def2 slen_scons_eq_rev) apply (tactic \step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1\) apply (tactic \step_tac (put_claset HOL_cs @{context} addSEs @{thms DefE}) 1\) @@ -201,7 +201,7 @@ done lemma fsfilter_fscons: - "A(C)x~> xs = (if x:A then x~> (A(C)xs) else A(C)xs)" + "A(C)x~> xs = (if x\A then x~> (A(C)xs) else A(C)xs)" apply (unfold fsfilter_def) apply (simp add: fscons_def2 If_and_if) done @@ -237,10 +237,10 @@ done lemma fstream_lub_lemma: - "\chain Y; (\i. Y i) = a\s\ \ (\j t. Y j = a\t) & (\X. chain X & (!i. ? j. Y j = a\X i) & (\i. X i) = s)" + "\chain Y; (\i. Y i) = a\s\ \ (\j t. Y j = a\t) \ (\X. chain X \ (\i. \j. Y j = a\X i) \ (\i. X i) = s)" apply (frule (1) fstream_lub_lemma1) apply (clarsimp) -apply (rule_tac x="%i. rt\(Y(i+j))" in exI) +apply (rule_tac x="\i. rt\(Y(i+j))" in exI) apply (rule conjI) apply (erule chain_shift [THEN chain_monofun]) apply safe diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/FOCUS/Stream_adm.thy --- a/src/HOL/HOLCF/FOCUS/Stream_adm.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/FOCUS/Stream_adm.thy Thu Feb 15 13:04:36 2018 +0100 @@ -56,15 +56,15 @@ lemma flatstream_adm_lemma: assumes 1: "Porder.chain Y" - assumes 2: "!i. P (Y i)" - assumes 3: "(!!Y. [| Porder.chain Y; !i. P (Y i); !k. ? j. enat k < #((Y j)::'a::flat stream)|] + assumes 2: "\i. P (Y i)" + assumes 3: "(\Y. [| Porder.chain Y; \i. P (Y i); \k. \j. enat k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))" shows "P(LUB i. Y i)" apply (rule increasing_chain_adm_lemma [OF 1 2]) apply (erule 3, assumption) apply (erule thin_rl) apply (rule allI) -apply (case_tac "!j. stream_finite (Y j)") +apply (case_tac "\j. stream_finite (Y j)") apply ( rule chain_incr) apply ( rule allI) apply ( drule spec) @@ -78,8 +78,8 @@ done (* should be without reference to stream length? *) -lemma flatstream_admI: "[|(!!Y. [| Porder.chain Y; !i. P (Y i); - !k. ? j. enat k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P" +lemma flatstream_admI: "[|(\Y. [| Porder.chain Y; \i. P (Y i); + \k. \j. enat k < #((Y j)::'a::flat stream)|] ==> P(LUB i. Y i))|]==> adm P" apply (unfold adm_def) apply (intro strip) apply (erule (1) flatstream_adm_lemma) @@ -92,8 +92,8 @@ by (rule order_trans) auto lemma stream_monoP2I: -"!!X. stream_monoP F ==> !i. ? l. !x y. - enat l <= #x --> (x::'a::flat stream) << y --> x:(F ^^ i) top --> y:(F ^^ i) top" +"\X. stream_monoP F \ \i. \l. \x y. + enat l \ #x \ (x::'a::flat stream) << y --> x \ (F ^^ i) top \ y \ (F ^^ i) top" apply (unfold stream_monoP_def) apply (safe) apply (rule_tac x="i*ia" in exI) @@ -119,9 +119,9 @@ apply (assumption) done -lemma stream_monoP2_gfp_admI: "[| !i. ? l. !x y. - enat l <= #x --> (x::'a::flat stream) << y --> x:(F ^^ i) top --> y:(F ^^ i) top; - inf_continuous F |] ==> adm (%x. x:gfp F)" +lemma stream_monoP2_gfp_admI: "[| \i. \l. \x y. + enat l \ #x \ (x::'a::flat stream) << y \ x \ (F ^^ i) top \ y \ (F ^^ i) top; + inf_continuous F |] ==> adm (\x. x \ gfp F)" apply (erule inf_continuous_gfp[of F, THEN ssubst]) apply (simp (no_asm)) apply (rule adm_lemmas) @@ -143,8 +143,8 @@ lemmas fstream_gfp_admI = stream_monoP2I [THEN stream_monoP2_gfp_admI] lemma stream_antiP2I: -"!!X. [|stream_antiP (F::(('a::flat stream)set => ('a stream set)))|] - ==> !i x y. x << y --> y:(F ^^ i) top --> x:(F ^^ i) top" +"\X. [|stream_antiP (F::(('a::flat stream)set => ('a stream set)))|] + ==> \i x y. x << y \ y \ (F ^^ i) top \ x \ (F ^^ i) top" apply (unfold stream_antiP_def) apply (rule allI) apply (induct_tac "i") @@ -170,8 +170,8 @@ done lemma stream_antiP2_non_gfp_admI: -"!!X. [|!i x y. x << y --> y:(F ^^ i) top --> x:(F ^^ i) top; inf_continuous F |] - ==> adm (%u. ~ u:gfp F)" +"\X. [|\i x y. x << y \ y \ (F ^^ i) top \ x \ (F ^^ i) top; inf_continuous F |] + ==> adm (\u. \ u \ gfp F)" apply (unfold adm_def) apply (simp add: inf_continuous_gfp) apply (fast dest!: is_ub_thelub) @@ -185,17 +185,17 @@ section "antitonP" -lemma antitonPD: "[| antitonP P; y:P; x< x:P" +lemma antitonPD: "[| antitonP P; y \ P; x< x \ P" apply (unfold antitonP_def) apply auto done -lemma antitonPI: "!x y. y:P --> x< x:P ==> antitonP P" +lemma antitonPI: "\x y. y \ P \ x< x \ P \ antitonP P" apply (unfold antitonP_def) apply (fast) done -lemma antitonP_adm_non_P: "antitonP P ==> adm (%u. u~:P)" +lemma antitonP_adm_non_P: "antitonP P \ adm (\u. u \ P)" apply (unfold adm_def) apply (auto dest: antitonPD elim: is_ub_thelub) done @@ -210,7 +210,7 @@ done lemma adm_set: -"{\i. Y i |Y. Porder.chain Y & (\i. Y i \ P)} \ P \ adm (\x. x\P)" +"{\i. Y i |Y. Porder.chain Y \ (\i. Y i \ P)} \ P \ adm (\x. x\P)" apply (unfold adm_def) apply (fast) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IMP/Denotational.thy --- a/src/HOL/HOLCF/IMP/Denotational.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IMP/Denotational.thy Thu Feb 15 13:04:36 2018 +0100 @@ -11,9 +11,9 @@ definition dlift :: "(('a::type) discr -> 'b::pcpo) => ('a lift -> 'b)" where - "dlift f = (LAM x. case x of UU => UU | Def y => f\(Discr y))" + "dlift f = (LAM x. case x of UU \ UU | Def y \ f\(Discr y))" -primrec D :: "com => state discr -> state lift" +primrec D :: "com \ state discr \ state lift" where "D(SKIP) = (LAM s. Def(undiscr s))" | "D(X ::= a) = (LAM s. Def((undiscr s)(X := aval a (undiscr s))))" @@ -37,7 +37,7 @@ "(dlift f\l = Def y) = (\x. l = Def x \ f\(Discr x) = Def y)" by (simp add: dlift_def split: lift.split) -lemma eval_implies_D: "(c,s) \ t ==> D c\(Discr s) = (Def t)" +lemma eval_implies_D: "(c,s) \ t \ D c\(Discr s) = (Def t)" apply (induct rule: big_step_induct) apply (auto) apply (subst fix_eq) @@ -46,7 +46,7 @@ apply simp done -lemma D_implies_eval: "!s t. D c\(Discr s) = (Def t) --> (c,s) \ t" +lemma D_implies_eval: "\s t. D c\(Discr s) = (Def t) \ (c,s) \ t" apply (induct c) apply fastforce apply fastforce diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/ABP/Lemmas.thy --- a/src/HOL/HOLCF/IOA/ABP/Lemmas.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/ABP/Lemmas.thy Thu Feb 15 13:04:36 2018 +0100 @@ -8,37 +8,37 @@ subsection \Logic\ -lemma and_de_morgan_and_absorbe: "(~(A&B)) = ((~A)&B| ~B)" +lemma and_de_morgan_and_absorbe: "(\(A\B)) = ((\A)\B\ \B)" by blast -lemma bool_if_impl_or: "(if C then A else B) --> (A|B)" +lemma bool_if_impl_or: "(if C then A else B) \ (A\B)" by auto -lemma exis_elim: "(? x. x=P & Q(x)) = Q(P)" +lemma exis_elim: "(\x. x=P \ Q(x)) = Q(P)" by blast subsection \Sets\ lemma set_lemmas: - "f(x) : (UN x. {f(x)})" - "f x y : (UN x y. {f x y})" - "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})" - "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})" + "f(x) \ (\x. {f(x)})" + "f x y \ (\x y. {f x y})" + "\a. (\x. a \ f(x)) \ a \ (\x. {f(x)})" + "\a. (\x y. a \ f x y) ==> a \ (\x y. {f x y})" by auto text \2 Lemmas to add to \set_lemmas\, used also for action handling, namely for Intersections and the empty list (compatibility of IOA!).\ -lemma singleton_set: "(UN b.{x. x=f(b)})= (UN b.{f(b)})" +lemma singleton_set: "(\b.{x. x=f(b)}) = (\b.{f(b)})" by blast -lemma de_morgan: "((A|B)=False) = ((~A)&(~B))" +lemma de_morgan: "((A\B)=False) = ((\A)\(\B))" by blast subsection \Lists\ -lemma cons_not_nil: "l ~= [] --> (? x xs. l = (x#xs))" +lemma cons_not_nil: "l \ [] \ (\x xs. l = (x#xs))" by (induct l) simp_all end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/CompoExecs.thy --- a/src/HOL/HOLCF/IOA/CompoExecs.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/CompoExecs.thy Thu Feb 15 13:04:36 2018 +0100 @@ -213,8 +213,8 @@ theorem compositionality_ex: "ex \ executions (A \ B) \ - Filter_ex (asig_of A) (ProjA ex) : executions A \ - Filter_ex (asig_of B) (ProjB ex) : executions B \ + Filter_ex (asig_of A) (ProjA ex) \ executions A \ + Filter_ex (asig_of B) (ProjB ex) \ executions B \ stutter (asig_of A) (ProjA ex) \ stutter (asig_of B) (ProjB ex) \ Forall (\x. fst x \ act (A \ B)) (snd ex)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/CompoScheds.thy --- a/src/HOL/HOLCF/IOA/CompoScheds.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/CompoScheds.thy Thu Feb 15 13:04:36 2018 +0100 @@ -52,9 +52,9 @@ schA = fst SchedsA; sigA = snd SchedsA; schB = fst SchedsB; sigB = snd SchedsB in - ({sch. Filter (%a. a:actions sigA)\sch : schA} \ - {sch. Filter (%a. a:actions sigB)\sch : schB} \ - {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch}, + ({sch. Filter (\a. a\actions sigA)\sch \ schA} \ + {sch. Filter (\a. a\actions sigB)\sch \ schB} \ + {sch. Forall (\x. x\(actions sigA Un actions sigB)) sch}, asig_comp sigA sigB))" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/CompoTraces.thy --- a/src/HOL/HOLCF/IOA/CompoTraces.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/CompoTraces.thy Thu Feb 15 13:04:36 2018 +0100 @@ -185,7 +185,7 @@ apply simp apply (rule Forall_Conc_impl [THEN mp]) apply (simp add: intA_is_not_actB int_is_act) - apply (case_tac "a:act B") + apply (case_tac "a\act B") text \\a \ A\, \a \ B\\ apply simp apply (rule Forall_Conc_impl [THEN mp]) @@ -604,7 +604,7 @@ apply (frule_tac y = "y2" in sym [THEN eq_imp_below, THEN divide_Seq]) apply (erule conjE)+ text \assumption \schA\\ - apply (drule_tac x = "schA" and g = "Filter (%a. a:act A) \rs" in subst_lemma2) + apply (drule_tac x = "schA" and g = "Filter (\a. a\act A) \rs" in subst_lemma2) apply assumption apply (simp add: int_is_not_ext) @@ -841,10 +841,10 @@ apply (simp add: intA_is_not_actB int_is_not_ext) text \conclusions concerning \LastActExtsch\, cannot be rewritten, as \LastActExtsmalli\ are looping\ - apply (drule_tac sch = "schB" and P = "%a. a:int B" in LastActExtsmall1) + apply (drule_tac sch = "schB" and P = "\a. a\int B" in LastActExtsmall1) apply (frule_tac ?sch1.0 = "x1" in LastActExtsmall2) apply assumption - apply (drule_tac sch = "x2" and P = "%a. a:int A" in LastActExtsmall1) + apply (drule_tac sch = "x2" and P = "\a. a\int A" in LastActExtsmall1) text \assumption \Forall schA\, and \Forall schA\ are performed by \ForallTL\, \ForallDropwhile\\ apply (simp add: ForallTL ForallDropwhile) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Abschannel.thy --- a/src/HOL/HOLCF/IOA/NTP/Abschannel.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Abschannel.thy Thu Feb 15 13:04:36 2018 +0100 @@ -102,28 +102,28 @@ trans_of_def asig_projections lemma in_srch_asig: - "S_msg(m) ~: actions(srch_asig) & - R_msg(m) ~: actions(srch_asig) & - S_pkt(pkt) : actions(srch_asig) & - R_pkt(pkt) : actions(srch_asig) & - S_ack(b) ~: actions(srch_asig) & - R_ack(b) ~: actions(srch_asig) & - C_m_s ~: actions(srch_asig) & - C_m_r ~: actions(srch_asig) & - C_r_s ~: actions(srch_asig) & C_r_r(m) ~: actions(srch_asig)" + "S_msg(m) \ actions(srch_asig) \ + R_msg(m) \ actions(srch_asig) \ + S_pkt(pkt) \ actions(srch_asig) \ + R_pkt(pkt) \ actions(srch_asig) \ + S_ack(b) \ actions(srch_asig) \ + R_ack(b) \ actions(srch_asig) \ + C_m_s \ actions(srch_asig) \ + C_m_r \ actions(srch_asig) \ + C_r_s \ actions(srch_asig) & C_r_r(m) \ actions(srch_asig)" by (simp add: unfold_renaming) lemma in_rsch_asig: - "S_msg(m) ~: actions(rsch_asig) & - R_msg(m) ~: actions(rsch_asig) & - S_pkt(pkt) ~: actions(rsch_asig) & - R_pkt(pkt) ~: actions(rsch_asig) & - S_ack(b) : actions(rsch_asig) & - R_ack(b) : actions(rsch_asig) & - C_m_s ~: actions(rsch_asig) & - C_m_r ~: actions(rsch_asig) & - C_r_s ~: actions(rsch_asig) & - C_r_r(m) ~: actions(rsch_asig)" + "S_msg(m) \ actions(rsch_asig) \ + R_msg(m) \ actions(rsch_asig) \ + S_pkt(pkt) \ actions(rsch_asig) \ + R_pkt(pkt) \ actions(rsch_asig) \ + S_ack(b) \ actions(rsch_asig) \ + R_ack(b) \ actions(rsch_asig) \ + C_m_s \ actions(rsch_asig) \ + C_m_r \ actions(rsch_asig) \ + C_r_s \ actions(rsch_asig) \ + C_r_r(m) \ actions(rsch_asig)" by (simp add: unfold_renaming) lemma srch_ioa_thm: "srch_ioa = diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Correctness.thy --- a/src/HOL/HOLCF/IOA/NTP/Correctness.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Correctness.thy Thu Feb 15 13:04:36 2018 +0100 @@ -27,18 +27,18 @@ \ lemma externals_lemma: - "a:externals(asig_of(Automata.restrict impl_ioa (externals spec_sig))) = + "a\externals(asig_of(Automata.restrict impl_ioa (externals spec_sig))) = (case a of - S_msg(m) => True - | R_msg(m) => True - | S_pkt(pkt) => False - | R_pkt(pkt) => False - | S_ack(b) => False - | R_ack(b) => False - | C_m_s => False - | C_m_r => False - | C_r_s => False - | C_r_r(m) => False)" + S_msg(m) \ True + | R_msg(m) \ True + | S_pkt(pkt) \ False + | R_pkt(pkt) \ False + | S_ack(b) \ False + | R_ack(b) \ False + | C_m_s \ False + | C_m_r \ False + | C_r_s \ False + | C_r_r(m) \ False)" apply (simp (no_asm) add: externals_def restrict_def restrict_asig_def Spec.sig_def asig_projections) apply (induct_tac "a") diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Impl.thy --- a/src/HOL/HOLCF/IOA/NTP/Impl.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Impl.thy Thu Feb 15 13:04:36 2018 +0100 @@ -28,9 +28,9 @@ (* Lemma 5.1 *) definition - "inv1(s) == - (!b. count (rsent(rec s)) b = count (srcvd(sen s)) b + count (rsch s) b) - & (!b. count (ssent(sen s)) b + "inv1(s) \ + (\b. count (rsent(rec s)) b = count (srcvd(sen s)) b + count (rsch s) b) + \ (\b. count (ssent(sen s)) b = hdr_sum (rrcvd(rec s)) b + hdr_sum (srch s) b)" (* Lemma 5.2 *) @@ -48,12 +48,12 @@ (* Lemma 5.3 *) definition - "inv3(s) == + "inv3(s) \ rbit(rec(s)) = sbit(sen(s)) - --> (!m. sq(sen(s))=[] | m ~= hd(sq(sen(s))) - --> count (rrcvd(rec s)) (sbit(sen(s)),m) + \ (\m. sq(sen(s))=[] | m \ hd(sq(sen(s))) + \ count (rrcvd(rec s)) (sbit(sen(s)),m) + count (srch s) (sbit(sen(s)),m) - <= count (rsent(rec s)) (~sbit(sen s)))" + \ count (rsent(rec s)) (~sbit(sen s)))" (* Lemma 5.4 *) definition "inv4(s) == rbit(rec(s)) = (~sbit(sen(s))) --> sq(sen(s)) ~= []" @@ -85,10 +85,10 @@ by (simp_all add: sen_def rec_def srch_def rsch_def) lemma [simp]: - "a:actions(sender_asig) - | a:actions(receiver_asig) - | a:actions(srch_asig) - | a:actions(rsch_asig)" + "a\actions(sender_asig) + \ a\actions(receiver_asig) + \ a\actions(srch_asig) + \ a\actions(rsch_asig)" by (induct a) simp_all declare split_paired_All [simp del] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Lemmas.thy --- a/src/HOL/HOLCF/IOA/NTP/Lemmas.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Lemmas.thy Thu Feb 15 13:04:36 2018 +0100 @@ -8,17 +8,17 @@ subsubsection \Logic\ -lemma neg_flip: "(X = (~ Y)) = ((~X) = Y)" +lemma neg_flip: "(X = (\ Y)) = ((\X) = Y)" by blast subsection \Sets\ lemma set_lemmas: - "f(x) : (UN x. {f(x)})" - "f x y : (UN x y. {f x y})" - "!!a. (!x. a ~= f(x)) ==> a ~: (UN x. {f(x)})" - "!!a. (!x y. a ~= f x y) ==> a ~: (UN x y. {f x y})" + "f(x) \ (\x. {f(x)})" + "f x y \ (\x y. {f x y})" + "\a. (\x. a \ f(x)) \ a \ (\x. {f(x)})" + "\a. (\x y. a \ f x y) \ a \ (\x y. {f x y})" by auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Multiset.thy --- a/src/HOL/HOLCF/IOA/NTP/Multiset.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Multiset.thy Thu Feb 15 13:04:36 2018 +0100 @@ -60,7 +60,7 @@ apply simp done -lemma countm_props: "!!M. (!x. P(x) --> Q(x)) ==> (countm M P <= countm M Q)" +lemma countm_props: "\M. (\x. P(x) \ Q(x)) \ (countm M P \ countm M Q)" apply (rule_tac M = "M" in Multiset.induction) apply (simp (no_asm) add: Multiset.countm_empty_def) apply (simp (no_asm) add: Multiset.countm_nonempty_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Packet.thy --- a/src/HOL/HOLCF/IOA/NTP/Packet.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Packet.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,16 +10,16 @@ 'msg packet = "bool * 'msg" definition - hdr :: "'msg packet => bool" where - "hdr == fst" + hdr :: "'msg packet \ bool" where + "hdr \ fst" definition - msg :: "'msg packet => 'msg" where - "msg == snd" + msg :: "'msg packet \ 'msg" where + "msg \ snd" text \Instantiation of a tautology?\ -lemma eq_packet_imp_eq_hdr: "!x. x = packet --> hdr(x) = hdr(packet)" +lemma eq_packet_imp_eq_hdr: "\x. x = packet \ hdr(x) = hdr(packet)" by simp declare hdr_def [simp] msg_def [simp] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Receiver.thy --- a/src/HOL/HOLCF/IOA/NTP/Receiver.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Receiver.thy Thu Feb 15 13:04:36 2018 +0100 @@ -79,16 +79,16 @@ (receiver_asig, {([],{|},{|},False,False)}, receiver_trans,{},{})" lemma in_receiver_asig: - "S_msg(m) ~: actions(receiver_asig)" - "R_msg(m) : actions(receiver_asig)" - "S_pkt(pkt) ~: actions(receiver_asig)" - "R_pkt(pkt) : actions(receiver_asig)" - "S_ack(b) : actions(receiver_asig)" - "R_ack(b) ~: actions(receiver_asig)" - "C_m_s ~: actions(receiver_asig)" - "C_m_r : actions(receiver_asig)" - "C_r_s ~: actions(receiver_asig)" - "C_r_r(m) : actions(receiver_asig)" + "S_msg(m) \ actions(receiver_asig)" + "R_msg(m) \ actions(receiver_asig)" + "S_pkt(pkt) \ actions(receiver_asig)" + "R_pkt(pkt) \ actions(receiver_asig)" + "S_ack(b) \ actions(receiver_asig)" + "R_ack(b) \ actions(receiver_asig)" + "C_m_s \ actions(receiver_asig)" + "C_m_r \ actions(receiver_asig)" + "C_r_s \ actions(receiver_asig)" + "C_r_r(m) \ actions(receiver_asig)" by (simp_all add: receiver_asig_def actions_def asig_projections) lemmas receiver_projections = rq_def rsent_def rrcvd_def rbit_def rsending_def diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Sender.thy --- a/src/HOL/HOLCF/IOA/NTP/Sender.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Sender.thy Thu Feb 15 13:04:36 2018 +0100 @@ -31,41 +31,41 @@ t = snd(snd(tr)) in case fst(snd(tr)) of - S_msg(m) => sq(t)=sq(s)@[m] & - ssent(t)=ssent(s) & - srcvd(t)=srcvd(s) & - sbit(t)=sbit(s) & + S_msg(m) => sq(t)=sq(s)@[m] \ + ssent(t)=ssent(s) \ + srcvd(t)=srcvd(s) \ + sbit(t)=sbit(s) \ ssending(t)=ssending(s) | R_msg(m) => False | - S_pkt(pkt) => hdr(pkt) = sbit(s) & - (? Q. sq(s) = (msg(pkt)#Q)) & - sq(t) = sq(s) & - ssent(t) = addm (ssent s) (sbit s) & - srcvd(t) = srcvd(s) & - sbit(t) = sbit(s) & - ssending(s) & + S_pkt(pkt) => hdr(pkt) = sbit(s) \ + (\Q. sq(s) = (msg(pkt)#Q)) \ + sq(t) = sq(s) \ + ssent(t) = addm (ssent s) (sbit s) \ + srcvd(t) = srcvd(s) \ + sbit(t) = sbit(s) \ + ssending(s) \ ssending(t) | R_pkt(pkt) => False | S_ack(b) => False | - R_ack(b) => sq(t)=sq(s) & - ssent(t)=ssent(s) & - srcvd(t) = addm (srcvd s) b & - sbit(t)=sbit(s) & + R_ack(b) => sq(t)=sq(s) \ + ssent(t)=ssent(s) \ + srcvd(t) = addm (srcvd s) b \ + sbit(t)=sbit(s) \ ssending(t)=ssending(s) | - C_m_s => count (ssent s) (~sbit s) < count (srcvd s) (~sbit s) & - sq(t)=sq(s) & - ssent(t)=ssent(s) & - srcvd(t)=srcvd(s) & - sbit(t)=sbit(s) & - ssending(s) & + C_m_s => count (ssent s) (~sbit s) < count (srcvd s) (~sbit s) \ + sq(t)=sq(s) \ + ssent(t)=ssent(s) \ + srcvd(t)=srcvd(s) \ + sbit(t)=sbit(s) \ + ssending(s) \ ~ssending(t) | C_m_r => False | - C_r_s => count (ssent s) (sbit s) <= count (srcvd s) (~sbit s) & - sq(t)=tl(sq(s)) & - ssent(t)=ssent(s) & - srcvd(t)=srcvd(s) & - sbit(t) = (~sbit(s)) & - ~ssending(s) & + C_r_s => count (ssent s) (sbit s) <= count (srcvd s) (~sbit s) \ + sq(t)=tl(sq(s)) \ + ssent(t)=ssent(s) \ + srcvd(t)=srcvd(s) \ + sbit(t) = (~sbit(s)) \ + ~ssending(s) \ ssending(t) | C_r_r(m) => False}" @@ -74,17 +74,17 @@ "sender_ioa = (sender_asig, {([],{|},{|},False,True)}, sender_trans,{},{})" -lemma in_sender_asig: - "S_msg(m) : actions(sender_asig)" - "R_msg(m) ~: actions(sender_asig)" - "S_pkt(pkt) : actions(sender_asig)" - "R_pkt(pkt) ~: actions(sender_asig)" - "S_ack(b) ~: actions(sender_asig)" - "R_ack(b) : actions(sender_asig)" - "C_m_s : actions(sender_asig)" - "C_m_r ~: actions(sender_asig)" - "C_r_s : actions(sender_asig)" - "C_r_r(m) ~: actions(sender_asig)" +lemma in_sender_asig: + "S_msg(m) \ actions(sender_asig)" + "R_msg(m) \ actions(sender_asig)" + "S_pkt(pkt) \ actions(sender_asig)" + "R_pkt(pkt) \ actions(sender_asig)" + "S_ack(b) \ actions(sender_asig)" + "R_ack(b) \ actions(sender_asig)" + "C_m_s \ actions(sender_asig)" + "C_m_r \ actions(sender_asig)" + "C_r_s \ actions(sender_asig)" + "C_r_r(m) \ actions(sender_asig)" by (simp_all add: sender_asig_def actions_def asig_projections) lemmas sender_projections = sq_def ssent_def srcvd_def sbit_def ssending_def diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/NTP/Spec.thy --- a/src/HOL/HOLCF/IOA/NTP/Spec.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/NTP/Spec.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,8 +10,8 @@ definition spec_sig :: "'m action signature" where - sig_def: "spec_sig = (UN m.{S_msg(m)}, - UN m.{R_msg(m)}, + sig_def: "spec_sig = (\m.{S_msg(m)}, + \m.{R_msg(m)}, {})" definition @@ -22,16 +22,16 @@ in case fst(snd(tr)) of - S_msg(m) => t = s@[m] | - R_msg(m) => s = (m#t) | - S_pkt(pkt) => False | - R_pkt(pkt) => False | - S_ack(b) => False | - R_ack(b) => False | - C_m_s => False | - C_m_r => False | - C_r_s => False | - C_r_r(m) => False}" + S_msg(m) \ t = s@[m] | + R_msg(m) \ s = (m#t) | + S_pkt(pkt) \ False | + R_pkt(pkt) \ False | + S_ack(b) \ False | + R_ack(b) \ False | + C_m_s \ False | + C_m_r \ False | + C_r_s \ False | + C_r_r(m) \ False}" definition spec_ioa :: "('m action, 'm list)ioa" where diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/RefCorrectness.thy --- a/src/HOL/HOLCF/IOA/RefCorrectness.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/RefCorrectness.thy Thu Feb 15 13:04:36 2018 +0100 @@ -147,7 +147,7 @@ lemma move_subprop3: "is_ref_map f C A \ reachable C s \ (s, a, t) \ trans_of C \ - laststate (f s,@x. move A x (f s) a (f t)) = (f t)" + laststate (f s, SOME x. move A x (f s) a (f t)) = (f t)" apply (cut_tac move_is_move) defer apply assumption+ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/SimCorrectness.thy --- a/src/HOL/HOLCF/IOA/SimCorrectness.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/SimCorrectness.thy Thu Feb 15 13:04:36 2018 +0100 @@ -223,7 +223,7 @@ \ lemma simulation_starts: - "is_simulation R C A \ s:starts_of C \ + "is_simulation R C A \ s\starts_of C \ let S' = SOME s'. (s, s') \ R \ s' \ starts_of A in (s, S') \ R \ S' \ starts_of A" apply (simp add: is_simulation_def corresp_ex_sim_def Int_non_empty Image_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/Simulations.thy --- a/src/HOL/HOLCF/IOA/Simulations.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/Simulations.thy Thu Feb 15 13:04:36 2018 +0100 @@ -38,11 +38,11 @@ definition is_history_relation :: "('s1 \ 's2) set \ ('a, 's1) ioa \ ('a, 's2) ioa \ bool" where "is_history_relation R C A \ - is_simulation R C A \ is_ref_map (\x. (SOME y. (x, y) \ R^-1)) A C" + is_simulation R C A \ is_ref_map (\x. (SOME y. (x, y) \ R\)) A C" definition is_prophecy_relation :: "('s1 \ 's2) set \ ('a, 's1) ioa \ ('a, 's2) ioa \ bool" where "is_prophecy_relation R C A \ - is_backward_simulation R C A \ is_ref_map (\x. (SOME y. (x, y) \ R^-1)) A C" + is_backward_simulation R C A \ is_ref_map (\x. (SOME y. (x, y) \ R\)) A C" lemma set_non_empty: "A \ {} \ (\x. x \ A)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/Storage/Action.thy --- a/src/HOL/HOLCF/IOA/Storage/Action.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/Storage/Action.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,7 +10,7 @@ datatype action = New | Loc nat | Free nat -lemma [cong]: "!!x. x = y ==> case_action a b c x = case_action a b c y" +lemma [cong]: "\x. x = y \ case_action a b c x = case_action a b c y" by simp end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/Storage/Correctness.thy --- a/src/HOL/HOLCF/IOA/Storage/Correctness.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/Storage/Correctness.thy Thu Feb 15 13:04:36 2018 +0100 @@ -16,7 +16,7 @@ k = fst c; b = snd c; used = fst a; c = snd a in - (! l:used. l < k) & b=c}" + (\l\used. l < k) \ b=c}" declare split_paired_Ex [simp del] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/Storage/Impl.thy --- a/src/HOL/HOLCF/IOA/Storage/Impl.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/Storage/Impl.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,8 +10,8 @@ definition impl_sig :: "action signature" where - "impl_sig = (UN l.{Free l} Un {New}, - UN l.{Loc l}, + "impl_sig = (\l.{Free l} \ {New}, + \l.{Loc l}, {})" definition @@ -22,18 +22,18 @@ in case fst(snd(tr)) of - New => k' = k & b' | - Loc l => b & l= k & k'= (Suc k) & ~b' | - Free l => k'=k & b'=b}" + New \ k' = k \ b' | + Loc l \ b \ l= k \ k'= (Suc k) \ \b' | + Free l \ k'=k \ b'=b}" definition impl_ioa :: "(action, nat * bool)ioa" where "impl_ioa = (impl_sig, {(0,False)}, impl_trans,{},{})" lemma in_impl_asig: - "New : actions(impl_sig) & - Loc l : actions(impl_sig) & - Free l : actions(impl_sig) " + "New \ actions(impl_sig) \ + Loc l \ actions(impl_sig) \ + Free l \ actions(impl_sig) " by (simp add: impl_sig_def actions_def asig_projections) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/IOA/Storage/Spec.thy --- a/src/HOL/HOLCF/IOA/Storage/Spec.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/IOA/Storage/Spec.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,24 +10,24 @@ definition spec_sig :: "action signature" where - "spec_sig = (UN l.{Free l} Un {New}, - UN l.{Loc l}, + "spec_sig = (\l.{Free l} \ {New}, + \l.{Loc l}, {})" definition - spec_trans :: "(action, nat set * bool)transition set" where + spec_trans :: "(action, nat set \ bool)transition set" where "spec_trans = {tr. let s = fst(tr); used = fst s; c = snd s; t = snd(snd(tr)); used' = fst t; c' = snd t in case fst(snd(tr)) of - New => used' = used & c' | - Loc l => c & l~:used & used'= used Un {l} & ~c' | - Free l => used'=used - {l} & c'=c}" + New \ used' = used \ c' | + Loc l \ c \ l \ used \ used'= used \ {l} \ \c' | + Free l \ used'=used - {l} \ c'=c}" definition - spec_ioa :: "(action, nat set * bool)ioa" where + spec_ioa :: "(action, nat set \ bool)ioa" where "spec_ioa = (spec_sig, {({},False)}, spec_trans,{},{})" end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/Library/Stream.thy --- a/src/HOL/HOLCF/Library/Stream.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/Library/Stream.thy Thu Feb 15 13:04:36 2018 +0100 @@ -361,7 +361,7 @@ lemma slen_empty_eq: "(#x = 0) = (x = \)" by (cases x) auto -lemma slen_scons_eq: "(enat (Suc n) < #x) = (? a y. x = a && y \ a \ \ \ enat n < #y)" +lemma slen_scons_eq: "(enat (Suc n) < #x) = (\a y. x = a && y \ a \ \ \ enat n < #y)" apply (auto, case_tac "x=UU",auto) apply (drule stream_exhaust_eq [THEN iffD1], auto) apply (case_tac "#y") apply simp_all diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/Tools/Domain/domain_induction.ML --- a/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/Tools/Domain/domain_induction.ML Thu Feb 15 13:04:36 2018 +0100 @@ -91,7 +91,7 @@ (******************************************************************************) val case_UU_allI = - @{lemma "(!!x. x ~= UU ==> P x) ==> P UU ==> ALL x. P x" by metis} + @{lemma "(\x. x \ UU \ P x) \ P UU \ \x. P x" by metis} fun prove_induction (comp_dbind : binding) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/HOLCF/ex/Loop.thy --- a/src/HOL/HOLCF/ex/Loop.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/HOLCF/ex/Loop.thy Thu Feb 15 13:04:36 2018 +0100 @@ -75,7 +75,7 @@ (* properties of while and iterations *) (* ------------------------------------------------------------------------- *) -lemma loop_lemma1: "\EX y. b\y = FF; iterate k\(step\b\g)\x = UU\ +lemma loop_lemma1: "\\y. b\y = FF; iterate k\(step\b\g)\x = UU\ \ iterate(Suc k)\(step\b\g)\x = UU" apply (simp (no_asm)) apply (rule trans) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hilbert_Choice.thy Thu Feb 15 13:04:36 2018 +0100 @@ -144,7 +144,7 @@ lemma inv_f_f: "inj f \ inv f (f x) = x" by simp -lemma f_inv_into_f: "y : f`A \ f (inv_into A f y) = y" +lemma f_inv_into_f: "y \ f`A \ f (inv_into A f y) = y" by (simp add: inv_into_def) (fast intro: someI2) lemma inv_into_f_eq: "inj_on f A \ x \ A \ f x = y \ inv_into A f y = x" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Arith2.thy --- a/src/HOL/Hoare/Arith2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Arith2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -9,13 +9,13 @@ imports Main begin -definition cd :: "[nat, nat, nat] => bool" - where "cd x m n \ x dvd m & x dvd n" +definition cd :: "[nat, nat, nat] \ bool" + where "cd x m n \ x dvd m \ x dvd n" -definition gcd :: "[nat, nat] => nat" - where "gcd m n = (SOME x. cd x m n & (!y.(cd y m n) --> y<=x))" +definition gcd :: "[nat, nat] \ nat" + where "gcd m n = (SOME x. cd x m n & (\y.(cd y m n) \ y\x))" -primrec fac :: "nat => nat" +primrec fac :: "nat \ nat" where "fac 0 = Suc 0" | "fac (Suc n) = Suc n * fac n" @@ -23,7 +23,7 @@ subsubsection \cd\ -lemma cd_nnn: "0 cd n n n" +lemma cd_nnn: "0 cd n n n" apply (simp add: cd_def) done @@ -37,12 +37,12 @@ apply blast done -lemma cd_diff_l: "n<=m ==> cd x m n = cd x (m-n) n" +lemma cd_diff_l: "n\m \ cd x m n = cd x (m-n) n" apply (unfold cd_def) apply (fastforce dest: dvd_diffD) done -lemma cd_diff_r: "m<=n ==> cd x m n = cd x m (n-m)" +lemma cd_diff_r: "m\n \ cd x m n = cd x m (n-m)" apply (unfold cd_def) apply (fastforce dest: dvd_diffD) done @@ -50,7 +50,7 @@ subsubsection \gcd\ -lemma gcd_nnn: "0 n = gcd n n" +lemma gcd_nnn: "0 n = gcd n n" apply (unfold gcd_def) apply (frule cd_nnn) apply (rule some_equality [symmetric]) @@ -62,17 +62,17 @@ apply (simp add: gcd_def cd_swap) done -lemma gcd_diff_l: "n<=m ==> gcd m n = gcd (m-n) n" +lemma gcd_diff_l: "n\m \ gcd m n = gcd (m-n) n" apply (unfold gcd_def) - apply (subgoal_tac "n<=m ==> !x. cd x m n = cd x (m-n) n") + apply (subgoal_tac "n\m \ \x. cd x m n = cd x (m-n) n") apply simp apply (rule allI) apply (erule cd_diff_l) done -lemma gcd_diff_r: "m<=n ==> gcd m n = gcd m (n-m)" +lemma gcd_diff_r: "m\n \ gcd m n = gcd m (n-m)" apply (unfold gcd_def) - apply (subgoal_tac "m<=n ==> !x. cd x m n = cd x m (n-m) ") + apply (subgoal_tac "m\n \ \x. cd x m n = cd x m (n-m) ") apply simp apply (rule allI) apply (erule cd_diff_r) @@ -82,7 +82,7 @@ subsubsection \pow\ lemma sq_pow_div2 [simp]: - "m mod 2 = 0 ==> ((n::nat)*n)^(m div 2) = n^m" + "m mod 2 = 0 \ ((n::nat)*n)^(m div 2) = n^m" apply (simp add: power2_eq_square [symmetric] power_mult [symmetric] minus_mod_eq_mult_div [symmetric]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Examples.thy --- a/src/HOL/Hoare/Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,29 +12,29 @@ subsection \multiplication by successive addition\ lemma multiply_by_add: "VARS m s a b - {a=A & b=B} + {a=A \ b=B} m := 0; s := 0; - WHILE m~=a - INV {s=m*b & a=A & b=B} + WHILE m\a + INV {s=m*b \ a=A \ b=B} DO s := s+b; m := m+(1::nat) OD {s = A*B}" by vcg_simp lemma multiply_by_add_time: "VARS m s a b t - {a=A & b=B & t=0} + {a=A \ b=B \ t=0} m := 0; t := t+1; s := 0; t := t+1; - WHILE m~=a - INV {s=m*b & a=A & b=B & t = 2*m + 2} + WHILE m\a + INV {s=m*b \ a=A \ b=B \ t = 2*m + 2} DO s := s+b; t := t+1; m := m+(1::nat); t := t+1 OD {s = A*B \ t = 2*A + 2}" by vcg_simp lemma multiply_by_add2: "VARS M N P :: int - {m=M & n=N} + {m=M \ n=N} IF M < 0 THEN M := -M; N := -N ELSE SKIP FI; P := 0; WHILE 0 < M - INV {0 <= M & (EX p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N)} + INV {0 \ M \ (\p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N)} DO P := P+N; M := M - 1 OD {P = m*n}" apply vcg_simp @@ -42,11 +42,11 @@ done lemma multiply_by_add2_time: "VARS M N P t :: int - {m=M & n=N & t=0} + {m=M \ n=N \ t=0} IF M < 0 THEN M := -M; t := t+1; N := -N; t := t+1 ELSE SKIP FI; P := 0; t := t+1; WHILE 0 < M - INV {0 \ M & (EX p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N & t \ 0 & t \ 2*(p-M)+3)} + INV {0 \ M & (\p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N & t \ 0 & t \ 2*(p-M)+3)} DO P := P+N; t := t+1; M := M - 1; t := t+1 OD {P = m*n & t \ 2*abs m + 3}" apply vcg_simp @@ -283,11 +283,11 @@ lemma zero_search: "VARS A i {True} i := 0; - WHILE i < length A & A!i ~= key - INV {!j. j A!j ~= key} + WHILE i < length A & A!i \ key + INV {\j. j A!j \ key} DO i := i+1 OD {(i < length A --> A!i = key) & - (i = length A --> (!j. j < length A --> A!j ~= key))}" + (i = length A --> (\j. j < length A \ A!j \ key))}" apply vcg_simp apply(blast elim!: less_SucE) done @@ -295,11 +295,11 @@ lemma zero_search_time: "VARS A i t {t=0} i := 0; t := t+1; - WHILE i < length A & A!i ~= key - INV {(\j. j A!j ~= key) & i \ length A & t = i+1} + WHILE i < length A \ A!i \ key + INV {(\j. j A!j \ key) \ i \ length A \ t = i+1} DO i := i+1; t := t+1 OD - {(i < length A --> A!i = key) & - (i = length A --> (!j. j < length A --> A!j ~= key)) & t \ length A + 1}" + {(i < length A \ A!i = key) \ + (i = length A \ (\j. j < length A --> A!j \ key)) \ t \ length A + 1}" apply vcg_simp apply(blast elim!: less_SucE) done @@ -318,22 +318,22 @@ lemma Partition: -"[| leq == %A i. !k. k A!k <= pivot; - geq == %A i. !k. i pivot <= A!k |] ==> +"[| leq == \A i. \k. k A!k \ pivot; + geq == \A i. \k. i k pivot \ A!k |] ==> VARS A u l {0 < length(A::('a::order)list)} l := 0; u := length A - Suc 0; - WHILE l <= u - INV {leq A l & geq A u & u u + INV {leq A l \ geq A u \ u l\length A} + DO WHILE l < length A \ A!l \ pivot + INV {leq A l & geq A u \ u l\length A} DO l := l+1 OD; - WHILE 0 < u & pivot <= A!u - INV {leq A l & geq A u & u A!u + INV {leq A l & geq A u \ u l\length A} DO u := u - 1 OD; - IF l <= u THEN A := A[l := A!u, u := A!l] ELSE SKIP FI + IF l \ u THEN A := A[l := A!u, u := A!l] ELSE SKIP FI OD - {leq A u & (!k. u A!k = pivot) & geq A l}" + {leq A u & (\k. u k A!k = pivot) \ geq A l}" \ \expand and delete abbreviations first\ apply (simp) apply (erule thin_rl)+ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Heap.thy --- a/src/HOL/Hoare/Heap.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Heap.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,10 +12,10 @@ datatype 'a ref = Null | Ref 'a -lemma not_Null_eq [iff]: "(x ~= Null) = (EX y. x = Ref y)" +lemma not_Null_eq [iff]: "(x \ Null) = (\y. x = Ref y)" by (induct x) auto -lemma not_Ref_eq [iff]: "(ALL y. x ~= Ref y) = (x = Null)" +lemma not_Ref_eq [iff]: "(\y. x \ Ref y) = (x = Null)" by (induct x) auto primrec addr :: "'a ref \ 'a" where @@ -68,12 +68,12 @@ cycle.\ lemma neq_dP: "p \ q \ Path h p Ps q \ distinct Ps \ - EX a Qs. p = Ref a & Ps = a#Qs & a \ set Qs" + \a Qs. p = Ref a \ Ps = a#Qs \ a \ set Qs" by (case_tac Ps, auto) lemma neq_dP_disp: "\ p \ q; distPath h p Ps q \ \ - EX a Qs. p = Ref a \ Ps = a#Qs \ a \ set Qs" + \a Qs. p = Ref a \ Ps = a#Qs \ a \ set Qs" apply (simp only:distPath_def) by (case_tac Ps, auto) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Hoare_Logic.thy --- a/src/HOL/Hoare/Hoare_Logic.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Hoare_Logic.thy Thu Feb 15 13:04:36 2018 +0100 @@ -40,7 +40,7 @@ "Sem (IF b THEN c1 ELSE c2 FI) s s'" definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" - where "Valid p c q \ (!s s'. Sem c s s' --> s : p --> s' : q)" + where "Valid p c q \ (\s s'. Sem c s s' \ s \ p \ s' \ q)" syntax diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Hoare_Logic_Abort.thy --- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Thu Feb 15 13:04:36 2018 +0100 @@ -42,7 +42,7 @@ "Sem (IF b THEN c1 ELSE c2 FI) s s'" definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" where - "Valid p c q == \s s'. Sem c s s' \ s : Some ` p \ s' : Some ` q" + "Valid p c q \ \s s'. Sem c s s' \ s \ Some ` p \ s' \ Some ` q" syntax @@ -101,7 +101,7 @@ subsection \Derivation of the proof rules and, most importantly, the VCG tactic\ -lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}" +lemma Compl_Collect: "-(Collect b) = {x. \(b x)}" by blast ML_file "hoare_tac.ML" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Pointer_Examples.thy --- a/src/HOL/Hoare/Pointer_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Pointer_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -148,9 +148,9 @@ first version uses a relation on @{typ"'a ref"}:\ lemma "VARS tl p - {(p,X) \ {(Ref x,tl x) |x. True}^*} + {(p,X) \ {(Ref x,tl x) |x. True}\<^sup>*} WHILE p \ Null \ p \ X - INV {(p,X) \ {(Ref x,tl x) |x. True}^*} + INV {(p,X) \ {(Ref x,tl x) |x. True}\<^sup>*} DO p := p^.tl OD {p = X}" apply vcg_simp @@ -164,9 +164,9 @@ text\Finally, a version based on a relation on type @{typ 'a}:\ lemma "VARS tl p - {p \ Null \ (addr p,X) \ {(x,y). tl x = Ref y}^*} + {p \ Null \ (addr p,X) \ {(x,y). tl x = Ref y}\<^sup>*} WHILE p \ Null \ p \ Ref X - INV {p \ Null \ (addr p,X) \ {(x,y). tl x = Ref y}^*} + INV {p \ Null \ (addr p,X) \ {(x,y). tl x = Ref y}\<^sup>*} DO p := p^.tl OD {p = Ref X}" apply vcg_simp @@ -241,7 +241,7 @@ THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI; s := r; WHILE p \ Null \ q \ Null - INV {EX rs ps qs a. Path tl r rs s \ List tl p ps \ List tl q qs \ + INV {\rs ps qs a. Path tl r rs s \ List tl p ps \ List tl q qs \ distinct(a # ps @ qs @ rs) \ s = Ref a \ merge(Ps,Qs,\x y. hd x \ hd y) = rs @ a # merge(ps,qs,\x y. hd x \ hd y) \ @@ -367,15 +367,15 @@ by (rule unproven) lemma "VARS hd tl p q r s - {islist tl p & Ps = list tl p \ islist tl q & Qs = list tl q \ + {islist tl p \ Ps = list tl p \ islist tl q \ Qs = list tl q \ set Ps \ set Qs = {} \ (p \ Null \ q \ Null)} IF cor (q = Null) (cand (p \ Null) (p^.hd \ q^.hd)) THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI; s := r; WHILE p \ Null \ q \ Null - INV {EX rs ps qs a. ispath tl r s & rs = path tl r s \ - islist tl p & ps = list tl p \ islist tl q & qs = list tl q \ + INV {\rs ps qs a. ispath tl r s \ rs = path tl r s \ + islist tl p \ ps = list tl p \ islist tl q \ qs = list tl q \ distinct(a # ps @ qs @ rs) \ s = Ref a \ merge(Ps,Qs,\x y. hd x \ hd y) = rs @ a # merge(ps,qs,\x y. hd x \ hd y) \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare/Pointers0.thy --- a/src/HOL/Hoare/Pointers0.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare/Pointers0.thy Thu Feb 15 13:04:36 2018 +0100 @@ -300,9 +300,9 @@ suffices to talk about reachability, i.e.\ we can use relations directly.\ lemma "VARS tl p - {(p,X) \ {(x,y). y = tl x & x \ Null}^*} + {(p,X) \ {(x,y). y = tl x & x \ Null}\<^sup>*} WHILE p \ Null \ p \ X - INV {(p,X) \ {(x,y). y = tl x & x \ Null}^*} + INV {(p,X) \ {(x,y). y = tl x & x \ Null}\<^sup>*} DO p := p^.tl OD {p = X}" apply vcg_simp @@ -337,7 +337,7 @@ THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI; s := r; WHILE p \ Null \ q \ Null - INV {EX rs ps qs. Path tl r rs s \ List tl p ps \ List tl q qs \ + INV {\rs ps qs. Path tl r rs s \ List tl p ps \ List tl q qs \ distinct(s # ps @ qs @ rs) \ s \ Null \ merge(Ps,Qs,\x y. hd x \ hd y) = rs @ s # merge(ps,qs,\x y. hd x \ hd y) \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Hoare_Parallel/RG_Tran.thy --- a/src/HOL/Hoare_Parallel/RG_Tran.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Thu Feb 15 13:04:36 2018 +0100 @@ -28,7 +28,7 @@ and ctrans :: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) where "P -c\ Q \ (P,Q) \ ctran" -| "P -c*\ Q \ (P,Q) \ ctran^*" +| "P -c*\ Q \ (P,Q) \ ctran\<^sup>*" | Basic: "(Some(Basic f), s) -c\ (None, f s)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/ACom.thy --- a/src/HOL/IMP/ACom.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/ACom.thy Thu Feb 15 13:04:36 2018 +0100 @@ -133,25 +133,25 @@ done lemma strip_eq_SKIP: - "strip C = SKIP \ (EX P. C = SKIP {P})" + "strip C = SKIP \ (\P. C = SKIP {P})" by (cases C) simp_all lemma strip_eq_Assign: - "strip C = x::=e \ (EX P. C = x::=e {P})" + "strip C = x::=e \ (\P. C = x::=e {P})" by (cases C) simp_all lemma strip_eq_Seq: - "strip C = c1;;c2 \ (EX C1 C2. C = C1;;C2 & strip C1 = c1 & strip C2 = c2)" + "strip C = c1;;c2 \ (\C1 C2. C = C1;;C2 & strip C1 = c1 & strip C2 = c2)" by (cases C) simp_all lemma strip_eq_If: "strip C = IF b THEN c1 ELSE c2 \ - (EX P1 P2 C1 C2 Q. C = IF b THEN {P1} C1 ELSE {P2} C2 {Q} & strip C1 = c1 & strip C2 = c2)" + (\P1 P2 C1 C2 Q. C = IF b THEN {P1} C1 ELSE {P2} C2 {Q} & strip C1 = c1 & strip C2 = c2)" by (cases C) simp_all lemma strip_eq_While: "strip C = WHILE b DO c1 \ - (EX I P C1 Q. C = {I} WHILE b DO {P} C1 {Q} & strip C1 = c1)" + (\I P C1 Q. C = {I} WHILE b DO {P} C1 {Q} & strip C1 = c1)" by (cases C) simp_all lemma [simp]: "shift (\p. a) n = (\p. a)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_Int0.thy Thu Feb 15 13:04:36 2018 +0100 @@ -212,14 +212,14 @@ text\Correctness:\ -lemma aval'_correct: "s : \\<^sub>s S \ aval a s : \(aval' a S)" +lemma aval'_correct: "s \ \\<^sub>s S \ aval a s \ \(aval' a S)" by (induct a) (auto simp: gamma_num' gamma_plus' \_fun_def) -lemma in_gamma_update: "\ s : \\<^sub>s S; i : \ a \ \ s(x := i) : \\<^sub>s(S(x := a))" +lemma in_gamma_update: "\ s \ \\<^sub>s S; i \ \ a \ \ s(x := i) \ \\<^sub>s(S(x := a))" by(simp add: \_fun_def) lemma gamma_Step_subcomm: - assumes "!!x e S. f1 x e (\\<^sub>o S) \ \\<^sub>o (f2 x e S)" "!!b S. g1 b (\\<^sub>o S) \ \\<^sub>o (g2 b S)" + assumes "\x e S. f1 x e (\\<^sub>o S) \ \\<^sub>o (f2 x e S)" "\b S. g1 b (\\<^sub>o S) \ \\<^sub>o (g2 b S)" shows "Step f1 g1 (\\<^sub>o S) (\\<^sub>c C) \ \\<^sub>c (Step f2 g2 S C)" by (induction C arbitrary: S) (auto simp: mono_gamma_o assms) @@ -406,7 +406,7 @@ shows "(\x\X. m (S2 x)) < (\x\X. m (S1 x))" proof- from assms(3) have 1: "\x\X. m(S1 x) \ m(S2 x)" by (simp add: m1) - from assms(2,3,4) have "EX x:X. S1 x < S2 x" + from assms(2,3,4) have "\x\X. S1 x < S2 x" by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans) hence 2: "\x\X. m(S1 x) > m(S2 x)" by (metis m2) from sum_strict_mono_ex1[OF \finite X\ 1 2] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_Int1.thy Thu Feb 15 13:04:36 2018 +0100 @@ -16,7 +16,7 @@ "aval' (V x) S = fun S x" | "aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" -lemma aval'_correct: "s : \\<^sub>s S \ aval a s : \(aval' a S)" +lemma aval'_correct: "s \ \\<^sub>s S \ aval a s \ \(aval' a S)" by (induction a) (auto simp: gamma_num' gamma_plus' \_st_def) lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom" @@ -26,7 +26,7 @@ proof(induction C arbitrary: S) qed (auto simp: assms intro!: mono_gamma_o sup_ge1 sup_ge2) -lemma in_gamma_update: "\ s : \\<^sub>s S; i : \ a \ \ s(x := i) : \\<^sub>s(update S x a)" +lemma in_gamma_update: "\ s \ \\<^sub>s S; i \ \ a \ \ s(x := i) \ \\<^sub>s(update S x a)" by(simp add: \_st_def) end @@ -188,7 +188,7 @@ shows "(\x\X. m (S2 x)) < (\x\X. m (S1 x))" proof- from assms(3) have 1: "\x\X. m(S1 x) \ m(S2 x)" by (simp add: m1) - from assms(2,3,4) have "EX x:X. S1 x < S2 x" + from assms(2,3,4) have "\x\X. S1 x < S2 x" by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans) hence 2: "\x\X. m(S1 x) > m(S2 x)" by (metis m2) from sum_strict_mono_ex1[OF \finite X\ 1 2] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_Int2.thy --- a/src/HOL/IMP/Abs_Int2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_Int2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -35,7 +35,7 @@ and gamma_bot[simp]: "\ \ = {}" begin -lemma in_gamma_inf: "x : \ a1 \ x : \ a2 \ x : \(a1 \ a2)" +lemma in_gamma_inf: "x \ \ a1 \ x \ \ a2 \ x \ \(a1 \ a2)" by (metis IntI inter_gamma_subset_gamma_inf set_mp) lemma gamma_inf: "\(a1 \ a2) = \ a1 \ \ a2" @@ -50,11 +50,11 @@ fixes test_num' :: "val \ 'av \ bool" and inv_plus' :: "'av \ 'av \ 'av \ 'av * 'av" and inv_less' :: "bool \ 'av \ 'av \ 'av * 'av" -assumes test_num': "test_num' i a = (i : \ a)" +assumes test_num': "test_num' i a = (i \ \ a)" and inv_plus': "inv_plus' a a1 a2 = (a\<^sub>1',a\<^sub>2') \ - i1 : \ a1 \ i2 : \ a2 \ i1+i2 : \ a \ i1 : \ a\<^sub>1' \ i2 : \ a\<^sub>2'" + i1 \ \ a1 \ i2 \ \ a2 \ i1+i2 \ \ a \ i1 \ \ a\<^sub>1' \ i2 \ \ a\<^sub>2'" and inv_less': "inv_less' (i11',a\<^sub>2') \ - i1 : \ a1 \ i2 : \ a2 \ i1 : \ a\<^sub>1' \ i2 : \ a\<^sub>2'" + i1 \ \ a1 \ i2 \ \ a2 \ i1 \ \ a\<^sub>1' \ i2 \ \ a\<^sub>2'" locale Abs_Int_inv = Val_inv where \ = \ @@ -62,14 +62,14 @@ begin lemma in_gamma_sup_UpI: - "s : \\<^sub>o S1 \ s : \\<^sub>o S2 \ s : \\<^sub>o(S1 \ S2)" + "s \ \\<^sub>o S1 \ s \ \\<^sub>o S2 \ s \ \\<^sub>o(S1 \ S2)" by (metis (hide_lams, no_types) sup_ge1 sup_ge2 mono_gamma_o subsetD) fun aval'' :: "aexp \ 'av st option \ 'av" where "aval'' e None = \" | "aval'' e (Some S) = aval' e S" -lemma aval''_correct: "s : \\<^sub>o S \ aval a s : \(aval'' a S)" +lemma aval''_correct: "s \ \\<^sub>o S \ aval a s \ \(aval'' a S)" by(cases S)(auto simp add: aval'_correct split: option.splits) subsubsection "Backward analysis" @@ -103,16 +103,16 @@ (let (a1,a2) = inv_less' res (aval'' e1 S) (aval'' e2 S) in inv_aval' e1 a1 (inv_aval' e2 a2 S))" -lemma inv_aval'_correct: "s : \\<^sub>o S \ aval e s : \ a \ s : \\<^sub>o (inv_aval' e a S)" +lemma inv_aval'_correct: "s \ \\<^sub>o S \ aval e s \ \ a \ s \ \\<^sub>o (inv_aval' e a S)" proof(induction e arbitrary: a S) case N thus ?case by simp (metis test_num') next case (V x) - obtain S' where "S = Some S'" and "s : \\<^sub>s S'" using \s : \\<^sub>o S\ + obtain S' where "S = Some S'" and "s \ \\<^sub>s S'" using \s \ \\<^sub>o S\ by(auto simp: in_gamma_option_iff) - moreover hence "s x : \ (fun S' x)" + moreover hence "s x \ \ (fun S' x)" by(simp add: \_st_def) - moreover have "s x : \ a" using V(2) by simp + moreover have "s x \ \ a" using V(2) by simp ultimately show ?case by(simp add: Let_def \_st_def) (metis mono_gamma emptyE in_gamma_inf gamma_bot subset_empty) @@ -122,7 +122,7 @@ by (auto split: prod.split) qed -lemma inv_bval'_correct: "s : \\<^sub>o S \ bv = bval b s \ s : \\<^sub>o(inv_bval' b bv S)" +lemma inv_bval'_correct: "s \ \\<^sub>o S \ bv = bval b s \ s \ \\<^sub>o(inv_bval' b bv S)" proof(induction b arbitrary: S bv) case Bc thus ?case by simp next diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_Int2_ivl.thy --- a/src/HOL/IMP/Abs_Int2_ivl.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_Int2_ivl.thy Thu Feb 15 13:04:36 2018 +0100 @@ -229,7 +229,7 @@ definition uminus_rep :: "eint2 \ eint2" where "uminus_rep p = (let (l,h) = p in (-h, -l))" -lemma \_uminus_rep: "i : \_rep p \ -i \ \_rep(uminus_rep p)" +lemma \_uminus_rep: "i \ \_rep p \ -i \ \_rep(uminus_rep p)" by(auto simp: uminus_rep_def \_rep_def image_def uminus_le_Fin_iff Fin_uminus_le_iff split: prod.split) @@ -240,7 +240,7 @@ instance .. end -lemma \_uminus: "i : \_ivl iv \ -i \ \_ivl(- iv)" +lemma \_uminus: "i \ \_ivl iv \ -i \ \_ivl(- iv)" by transfer (rule \_uminus_rep) lemma uminus_nice: "-[l,h] = [-h,-l]" @@ -276,7 +276,7 @@ (auto simp add: above_rep_def \_rep_cases is_empty_rep_def split: extended.splits) -lemma \_belowI: "i : \_ivl iv \ j \ i \ j : \_ivl(below iv)" +lemma \_belowI: "i \ \_ivl iv \ j \ i \ j \ \_ivl(below iv)" by transfer (auto simp add: below_rep_def \_rep_cases is_empty_rep_def split: extended.splits) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_Int3.thy --- a/src/HOL/IMP/Abs_Int3.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_Int3.thy Thu Feb 15 13:04:36 2018 +0100 @@ -466,7 +466,7 @@ assumes P_f: "\C. P C \ P(f C)" and P_widen: "\C1 C2. P C1 \ P C2 \ P(C1 \ C2)" and m_widen: "\C1 C2. P C1 \ P C2 \ ~ C2 \ C1 \ m(C1 \ C2) < m C1" -and "P C" shows "EX C'. iter_widen f C = Some C'" +and "P C" shows "\C'. iter_widen f C = Some C'" proof(simp add: iter_widen_def, rule measure_while_option_Some[where P = P and f=m]) show "P C" by(rule \P C\) @@ -481,7 +481,7 @@ and P_narrow: "\C1 C2. P C1 \ P C2 \ P(C1 \ C2)" and mono: "\C1 C2. P C1 \ P C2 \ C1 \ C2 \ f C1 \ f C2" and n_narrow: "\C1 C2. P C1 \ P C2 \ C2 \ C1 \ C1 \ C2 < C1 \ n(C1 \ C2) < n C1" -and init: "P C" "f C \ C" shows "EX C'. iter_narrow f C = Some C'" +and init: "P C" "f C \ C" shows "\C'. iter_narrow f C = Some C'" proof(simp add: iter_narrow_def, rule measure_while_option_Some[where f=n and P = "%C. P C \ f C \ C"]) show "P C \ f C \ C" using init by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Abs_State.thy --- a/src/HOL/IMP/Abs_State.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Abs_State.thy Thu Feb 15 13:04:36 2018 +0100 @@ -157,7 +157,7 @@ by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2]) lemma in_gamma_option_iff: - "x : \_option r u \ (\u'. u = Some u' \ x : r u')" + "x \ \_option r u \ (\u'. u = Some u' \ x \ r u')" by (cases u) auto end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Collecting.thy --- a/src/HOL/IMP/Collecting.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Collecting.thy Thu Feb 15 13:04:36 2018 +0100 @@ -109,7 +109,7 @@ subsubsection "Collecting semantics" -definition "step = Step (\x e S. {s(x := aval e s) |s. s : S}) (\b S. {s:S. bval b s})" +definition "step = Step (\x e S. {s(x := aval e s) |s. s \ S}) (\b S. {s:S. bval b s})" definition CS :: "com \ state set acom" where "CS c = lfp c (step UNIV)" @@ -195,7 +195,7 @@ by (auto simp: step_def post_def) have "step {s \ I. bval b s} C' \ C'" by (rule order_trans[OF mono2_step[OF order_refl \{s \ I. bval b s} \ P\] \step P C' \ C'\]) - have "s1: {s:I. bval b s}" using \s1 \ S\ \S \ I\ \bval b s1\ by auto + have "s1 \ {s\I. bval b s}" using \s1 \ S\ \S \ I\ \bval b s1\ by auto note s2_in_post_C' = WhileTrue.IH(1)[OF \strip C' = c'\ this \step {s \ I. bval b s} C' \ C'\] from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' \step (post C') C \ C\] show ?case . @@ -207,7 +207,7 @@ lemma big_step_lfp: "\ (c,s) \ t; s \ S \ \ t \ post(lfp c (step S))" by(auto simp add: post_lfp intro: big_step_post_step) -lemma big_step_CS: "(c,s) \ t \ t : post(CS c)" +lemma big_step_CS: "(c,s) \ t \ t \ post(CS c)" by(simp add: CS_def big_step_lfp) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Collecting1.thy --- a/src/HOL/IMP/Collecting1.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Collecting1.thy Thu Feb 15 13:04:36 2018 +0100 @@ -27,12 +27,12 @@ definition steps :: "state \ com \ nat \ state set acom" where "steps s c n = ((step {})^^n) (step {s} (annotate (\p. {}) c))" -lemma steps_approx_fix_step: assumes "step S cs = cs" and "s:S" +lemma steps_approx_fix_step: assumes "step S cs = cs" and "s \ S" shows "steps s (strip cs) n \ cs" proof- let ?bot = "annotate (\p. {}) (strip cs)" have "?bot \ cs" by(induction cs) auto - from step_preserves_le[OF assms(1)_ this, of "{s}"] \s:S\ + from step_preserves_le[OF assms(1)_ this, of "{s}"] \s \ S\ have 1: "step {s} ?bot \ cs" by simp from steps_empty_preserves_le[OF assms(1) 1] show ?thesis by(simp add: steps_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Complete_Lattice.thy --- a/src/HOL/IMP/Complete_Lattice.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Complete_Lattice.thy Thu Feb 15 13:04:36 2018 +0100 @@ -7,25 +7,25 @@ locale Complete_Lattice = fixes L :: "'a::order set" and Glb :: "'a set \ 'a" assumes Glb_lower: "A \ L \ a \ A \ Glb A \ a" -and Glb_greatest: "b : L \ \a\A. b \ a \ b \ Glb A" -and Glb_in_L: "A \ L \ Glb A : L" +and Glb_greatest: "b \ L \ \a\A. b \ a \ b \ Glb A" +and Glb_in_L: "A \ L \ Glb A \ L" begin definition lfp :: "('a \ 'a) \ 'a" where "lfp f = Glb {a : L. f a \ a}" -lemma index_lfp: "lfp f : L" +lemma index_lfp: "lfp f \ L" by(auto simp: lfp_def intro: Glb_in_L) lemma lfp_lowerbound: - "\ a : L; f a \ a \ \ lfp f \ a" + "\ a \ L; f a \ a \ \ lfp f \ a" by (auto simp add: lfp_def intro: Glb_lower) lemma lfp_greatest: - "\ a : L; \u. \ u : L; f u \ u\ \ a \ u \ \ a \ lfp f" + "\ a \ L; \u. \ u \ L; f u \ u\ \ a \ u \ \ a \ lfp f" by (auto simp add: lfp_def intro: Glb_greatest) -lemma lfp_unfold: assumes "\x. f x : L \ x : L" +lemma lfp_unfold: assumes "\x. f x \ L \ x \ L" and mono: "mono f" shows "lfp f = f (lfp f)" proof- note assms(1)[simp] index_lfp[simp] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Def_Init_Small.thy --- a/src/HOL/IMP/Def_Init_Small.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Def_Init_Small.thy Thu Feb 15 13:04:36 2018 +0100 @@ -28,7 +28,7 @@ subsection "Soundness wrt Small Steps" theorem progress: - "D (dom s) c A' \ c \ SKIP \ EX cs'. (c,s) \ cs'" + "D (dom s) c A' \ c \ SKIP \ \cs'. (c,s) \ cs'" proof (induction c arbitrary: s A') case Assign thus ?case by auto (metis aval_Some small_step.Assign) next @@ -38,7 +38,7 @@ by(cases bv)(auto intro: small_step.IfTrue small_step.IfFalse) qed (fastforce intro: small_step.intros)+ -lemma D_mono: "D A c M \ A \ A' \ EX M'. D A' c M' & M <= M'" +lemma D_mono: "D A c M \ A \ A' \ \M'. D A' c M' & M <= M'" proof (induction c arbitrary: A A' M) case Seq thus ?case by auto (metis D.intros(3)) next @@ -55,7 +55,7 @@ qed (auto intro: D.intros) theorem D_preservation: - "(c,s) \ (c',s') \ D (dom s) c A \ EX A'. D (dom s') c' A' & A <= A'" + "(c,s) \ (c',s') \ D (dom s) c A \ \A'. D (dom s') c' A' & A <= A'" proof (induction arbitrary: A rule: small_step_induct) case (While b c s) then obtain A' where A': "vars b \ dom s" "A = dom s" "D (dom s) c A'" by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Denotational.thy --- a/src/HOL/IMP/Denotational.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Denotational.thy Thu Feb 15 13:04:36 2018 +0100 @@ -44,7 +44,7 @@ abbreviation Big_step :: "com \ com_den" where "Big_step c \ {(s,t). (c,s) \ t}" -lemma Big_step_if_D: "(s,t) \ D(c) \ (s,t) : Big_step c" +lemma Big_step_if_D: "(s,t) \ D(c) \ (s,t) \ Big_step c" proof (induction c arbitrary: s t) case Seq thus ?case by fastforce next diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Fold.thy --- a/src/HOL/IMP/Fold.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Fold.thy Thu Feb 15 13:04:36 2018 +0100 @@ -11,7 +11,7 @@ "afold (Plus e1 e2) t = (case (afold e1 t, afold e2 t) of (N n1, N n2) \ N(n1+n2) | (e1',e2') \ Plus e1' e2')" -definition "approx t s \ (ALL x k. t x = Some k \ s x = k)" +definition "approx t s \ (\x k. t x = Some k \ s x = k)" theorem aval_afold[simp]: assumes "approx t s" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Live.thy --- a/src/HOL/IMP/Live.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Live.thy Thu Feb 15 13:04:36 2018 +0100 @@ -176,23 +176,23 @@ text\Now the opposite direction.\ lemma SKIP_bury[simp]: - "SKIP = bury c X \ c = SKIP | (EX x a. c = x::=a & x \ X)" + "SKIP = bury c X \ c = SKIP | (\x a. c = x::=a & x \ X)" by (cases c) auto -lemma Assign_bury[simp]: "x::=a = bury c X \ c = x::=a & x : X" +lemma Assign_bury[simp]: "x::=a = bury c X \ c = x::=a \ x \ X" by (cases c) auto lemma Seq_bury[simp]: "bc\<^sub>1;;bc\<^sub>2 = bury c X \ - (EX c\<^sub>1 c\<^sub>2. c = c\<^sub>1;;c\<^sub>2 & bc\<^sub>2 = bury c\<^sub>2 X & bc\<^sub>1 = bury c\<^sub>1 (L c\<^sub>2 X))" + (\c\<^sub>1 c\<^sub>2. c = c\<^sub>1;;c\<^sub>2 & bc\<^sub>2 = bury c\<^sub>2 X & bc\<^sub>1 = bury c\<^sub>1 (L c\<^sub>2 X))" by (cases c) auto lemma If_bury[simp]: "IF b THEN bc1 ELSE bc2 = bury c X \ - (EX c1 c2. c = IF b THEN c1 ELSE c2 & + (\c1 c2. c = IF b THEN c1 ELSE c2 & bc1 = bury c1 X & bc2 = bury c2 X)" by (cases c) auto lemma While_bury[simp]: "WHILE b DO bc' = bury c X \ - (EX c'. c = WHILE b DO c' & bc' = bury c' (L (WHILE b DO c') X))" + (\c'. c = WHILE b DO c' & bc' = bury c' (L (WHILE b DO c') X))" by (cases c) auto theorem bury_correct2: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/Small_Step.thy --- a/src/HOL/IMP/Small_Step.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/Small_Step.thy Thu Feb 15 13:04:36 2018 +0100 @@ -173,7 +173,7 @@ subsection "Final configurations and infinite reductions" -definition "final cs \ \(EX cs'. cs \ cs')" +definition "final cs \ \(\cs'. cs \ cs')" lemma finalD: "final (c,s) \ c = SKIP" apply(simp add: final_def) @@ -188,7 +188,7 @@ terminates:\ lemma big_iff_small_termination: - "(EX t. cs \ t) \ (EX cs'. cs \* cs' \ final cs')" + "(\t. cs \ t) \ (\cs'. cs \* cs' \ final cs')" by(simp add: big_iff_small final_iff_SKIP) text\This is the same as saying that the absence of a big step result is diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/VCG_Total_EX.thy --- a/src/HOL/IMP/VCG_Total_EX.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/VCG_Total_EX.thy Thu Feb 15 13:04:36 2018 +0100 @@ -36,7 +36,7 @@ "pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" | "pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (\s. if bval b s then pre C\<^sub>1 Q s else pre C\<^sub>2 Q s)" | -"pre ({I} WHILE b DO C) Q = (\s. EX n. I n s)" +"pre ({I} WHILE b DO C) Q = (\s. \n. I n s)" text\Verification condition:\ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMP/VCG_Total_EX2.thy --- a/src/HOL/IMP/VCG_Total_EX2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMP/VCG_Total_EX2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -42,7 +42,7 @@ "pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" | "pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (\l s. if bval b s then pre C\<^sub>1 Q l s else pre C\<^sub>2 Q l s)" | -"pre ({I/x} WHILE b DO C) Q = (\l s. EX n. I (l(x:=n)) s)" +"pre ({I/x} WHILE b DO C) Q = (\l s. \n. I (l(x:=n)) s)" text\Verification condition:\ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMPP/Com.thy --- a/src/HOL/IMPP/Com.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMPP/Com.thy Thu Feb 15 13:04:36 2018 +0100 @@ -80,7 +80,7 @@ definition WT_bodies :: bool where - "WT_bodies = (!(pn,b):set bodies. WT b)" + "WT_bodies = (\(pn,b) \ set bodies. WT b)" ML \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMPP/Hoare.thy --- a/src/HOL/IMPP/Hoare.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMPP/Hoare.thy Thu Feb 15 13:04:36 2018 +0100 @@ -34,14 +34,14 @@ definition triple_valid :: "nat => 'a triple => bool" ( "|=_:_" [0 , 58] 57) where "|=n:t = (case t of {P}.c.{Q} => - !Z s. P Z s --> (!s'. -n-> s' --> Q Z s'))" + \Z s. P Z s \ (\s'. -n-> s' \ Q Z s'))" abbreviation triples_valid :: "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57) where "||=n:G == Ball G (triple_valid n)" definition hoare_valids :: "'a triple set => 'a triple set => bool" ("_||=_" [58, 58] 57) where - "G||=ts = (!n. ||=n:G --> ||=n:ts)" + "G||=ts = (\n. ||=n:G --> ||=n:ts)" abbreviation hoare_valid :: "'a triple set => 'a triple => bool" ("_|=_" [58, 58] 57) where "G |=t == G||={t}" @@ -68,8 +68,8 @@ | weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts" -| conseq: "!Z s. P Z s --> (? P' Q'. G|-{P'}.c.{Q'} & - (!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s')) +| conseq: "\Z s. P Z s \ (\P' Q'. G|-{P'}.c.{Q'} \ + (\s'. (\Z'. P' Z' s \ Q' Z' s') \ Q Z s')) ==> G|-{P}.c.{Q}" @@ -108,7 +108,7 @@ section \Soundness and relative completeness of Hoare rules wrt operational semantics\ lemma single_stateE: - "state_not_singleton ==> !t. (!s::state. s = t) --> False" + "state_not_singleton \ \t. (\s::state. s = t) \ False" apply (unfold state_not_singleton_def) apply clarify apply (case_tac "ta = t") @@ -122,7 +122,7 @@ subsection "validity" lemma triple_valid_def2: - "|=n:{P}.c.{Q} = (!Z s. P Z s --> (!s'. -n-> s' --> Q Z s'))" + "|=n:{P}.c.{Q} = (\Z s. P Z s \ (\s'. -n-> s' \ Q Z s'))" apply (unfold triple_valid_def) apply auto done @@ -152,26 +152,26 @@ subsection "derived rules" -lemma conseq12: "[| G|-{P'}.c.{Q'}; !Z s. P Z s --> - (!s'. (!Z'. P' Z' s --> Q' Z' s') --> Q Z s') |] +lemma conseq12: "[| G|-{P'}.c.{Q'}; \Z s. P Z s \ + (\s'. (\Z'. P' Z' s \ Q' Z' s') --> Q Z s') |] ==> G|-{P}.c.{Q}" apply (rule hoare_derivs.conseq) apply blast done -lemma conseq1: "[| G|-{P'}.c.{Q}; !Z s. P Z s --> P' Z s |] ==> G|-{P}.c.{Q}" +lemma conseq1: "[| G|-{P'}.c.{Q}; \Z s. P Z s \ P' Z s |] ==> G|-{P}.c.{Q}" apply (erule conseq12) apply fast done -lemma conseq2: "[| G|-{P}.c.{Q'}; !Z s. Q' Z s --> Q Z s |] ==> G|-{P}.c.{Q}" +lemma conseq2: "[| G|-{P}.c.{Q'}; \Z s. Q' Z s \ Q Z s |] ==> G|-{P}.c.{Q}" apply (erule conseq12) apply fast done -lemma Body1: "[| G Un (%p. {P p}. BODY p .{Q p})`Procs - ||- (%p. {P p}. the (body p) .{Q p})`Procs; - pn:Procs |] ==> G|-{P pn}. BODY pn .{Q pn}" +lemma Body1: "[| G Un (\p. {P p}. BODY p .{Q p})`Procs + ||- (\p. {P p}. the (body p) .{Q p})`Procs; + pn \ Procs |] ==> G|-{P pn}. BODY pn .{Q pn}" apply (drule hoare_derivs.Body) apply (erule hoare_derivs.weaken) apply fast @@ -184,17 +184,17 @@ apply clarsimp done -lemma escape: "[| !Z s. P Z s --> G|-{%Z s'. s'=s}.c.{%Z'. Q Z} |] ==> G|-{P}.c.{Q}" +lemma escape: "[| \Z s. P Z s \ G|-{\Z s'. s'=s}.c.{\Z'. Q Z} |] ==> G|-{P}.c.{Q}" apply (rule hoare_derivs.conseq) apply fast done -lemma "constant": "[| C ==> G|-{P}.c.{Q} |] ==> G|-{%Z s. P Z s & C}.c.{Q}" +lemma "constant": "[| C ==> G|-{P}.c.{Q} |] ==> G|-{\Z s. P Z s & C}.c.{Q}" apply (rule hoare_derivs.conseq) apply fast done -lemma LoopF: "G|-{%Z s. P Z s & ~b s}.WHILE b DO c.{P}" +lemma LoopF: "G|-{\Z s. P Z s \ \b s}.WHILE b DO c.{P}" apply (rule hoare_derivs.Loop [THEN conseq2]) apply (simp_all (no_asm)) apply (rule hoare_derivs.conseq) @@ -207,7 +207,7 @@ by (etac hoare_derivs.asm 1); qed "thin"; *) -lemma thin [rule_format]: "G'||-ts ==> !G. G' <= G --> G||-ts" +lemma thin [rule_format]: "G'||-ts \ \G. G' <= G \ G||-ts" apply (erule hoare_derivs.induct) apply (tactic \ALLGOALS (EVERY'[clarify_tac @{context}, REPEAT o smp_tac @{context} 1])\) apply (rule hoare_derivs.empty) @@ -232,7 +232,7 @@ done lemma finite_pointwise [rule_format (no_asm)]: "[| finite U; - !p. G |- {P' p}.c0 p.{Q' p} --> G |- {P p}.c0 p.{Q p} |] ==> + \p. G |- {P' p}.c0 p.{Q' p} --> G |- {P p}.c0 p.{Q p} |] ==> G||-(%p. {P' p}.c0 p.{Q' p}) ` U --> G||-(%p. {P p}.c0 p.{Q p}) ` U" apply (erule finite_induct) apply simp @@ -251,7 +251,7 @@ apply (unfold hoare_valids_def) apply (simp (no_asm_use) add: triple_valid_def2) apply (rule allI) -apply (subgoal_tac "!d s s'. -n-> s' --> d = WHILE b DO c --> ||=n:G --> (!Z. P Z s --> P Z s' & ~b s') ") +apply (subgoal_tac "\d s s'. -n-> s' --> d = WHILE b DO c --> ||=n:G --> (\Z. P Z s --> P Z s' & ~b s') ") apply (erule thin_rl, fast) apply ((rule allI)+, rule impI) apply (erule evaln.induct) @@ -302,16 +302,16 @@ (* Both versions *) (*unused*) -lemma MGT_alternI: "G|-MGT c ==> - G|-{%Z s0. !s1. -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1}" +lemma MGT_alternI: "G|-MGT c \ + G|-{\Z s0. \s1. -c-> s1 \ Z=s1}. c .{\Z s1. Z=s1}" apply (unfold MGT_def) apply (erule conseq12) apply auto done (* requires com_det *) -lemma MGT_alternD: "state_not_singleton ==> - G|-{%Z s0. !s1. -c-> s1 --> Z=s1}. c .{%Z s1. Z=s1} ==> G|-MGT c" +lemma MGT_alternD: "state_not_singleton \ + G|-{\Z s0. \s1. -c-> s1 \ Z=s1}. c .{\Z s1. Z=s1} \ G|-MGT c" apply (unfold MGT_def) apply (erule conseq12) apply auto @@ -332,8 +332,8 @@ declare WTs_elim_cases [elim!] declare not_None_eq [iff] (* requires com_det, escape (i.e. hoare_derivs.conseq) *) -lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton ==> - !pn:dom body. G|-{=}.BODY pn.{->} ==> WT c --> G|-{=}.c.{->}" +lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton \ + \pn \ dom body. G|-{=}.BODY pn.{->} \ WT c --> G|-{=}.c.{->}" apply (induct_tac c) apply (tactic \ALLGOALS (clarsimp_tac @{context})\) prefer 7 apply (fast intro: domI) @@ -362,10 +362,10 @@ lemma nesting_lemma [rule_format]: assumes "!!G ts. ts <= G ==> P G ts" and "!!G pn. P (insert (mgt_call pn) G) {mgt(the(body pn))} ==> P G {mgt_call pn}" - and "!!G c. [| wt c; !pn:U. P G {mgt_call pn} |] ==> P G {mgt c}" - and "!!pn. pn : U ==> wt (the (body pn))" + and "!!G c. [| wt c; \pn\U. P G {mgt_call pn} |] ==> P G {mgt c}" + and "!!pn. pn \ U ==> wt (the (body pn))" shows "finite U ==> uG = mgt_call`U ==> - !G. G <= uG --> n <= card uG --> card G = card uG - n --> (!c. wt c --> P G {mgt c})" + \G. G <= uG --> n <= card uG --> card G = card uG - n --> (\c. wt c --> P G {mgt c})" apply (induct_tac n) apply (tactic \ALLGOALS (clarsimp_tac @{context})\) apply (subgoal_tac "G = mgt_call ` U") @@ -378,7 +378,7 @@ apply fast apply (erule assms(3-)) (*MGF_lemma1*) apply (rule ballI) -apply (case_tac "mgt_call pn : G") +apply (case_tac "mgt_call pn \ G") apply (rule assms) (*hoare_derivs.asm*) apply fast apply (rule assms(2-)) (*MGT_BodyN*) @@ -484,7 +484,7 @@ apply (fast elim: conseq12) done (* analogue conj non-derivable *) -lemma hoare_SkipI: "(!Z s. P Z s --> Q Z s) ==> G|-{P}. SKIP .{Q}" +lemma hoare_SkipI: "(\Z s. P Z s \ Q Z s) \ G|-{P}. SKIP .{Q}" apply (rule conseq12) apply (rule hoare_derivs.Skip) apply fast @@ -504,7 +504,7 @@ done -lemma weak_Local: "[| G|-{P}. c .{Q}; !k Z s. Q Z s --> Q Z (s[Loc Y::=k]) |] ==> +lemma weak_Local: "[| G|-{P}. c .{Q}; \k Z s. Q Z s --> Q Z (s[Loc Y::=k]) |] ==> G|-{%Z s. P Z (s[Loc Y::=a s])}. LOCAL Y:=a IN c .{Q}" apply (rule export_s) apply (rule hoare_derivs.Local) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMPP/Misc.thy --- a/src/HOL/IMPP/Misc.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMPP/Misc.thy Thu Feb 15 13:04:36 2018 +0100 @@ -81,7 +81,7 @@ (*unused*) lemma classic_Local_valid: -"!v. G|={%Z s. P Z (s[Loc Y::=v]) & s = a (s[Loc Y::=v])}. +"\v. G|={%Z s. P Z (s[Loc Y::=v]) & s = a (s[Loc Y::=v])}. c .{%Z s. Q Z (s[Loc Y::=v])} ==> G|={P}. LOCAL Y:=a IN c .{Q}" apply (unfold hoare_valids_def) apply (simp (no_asm_use) add: triple_valid_def2) @@ -96,7 +96,7 @@ apply simp done -lemma classic_Local: "!v. G|-{%Z s. P Z (s[Loc Y::=v]) & s = a (s[Loc Y::=v])}. +lemma classic_Local: "\v. G|-{%Z s. P Z (s[Loc Y::=v]) & s = a (s[Loc Y::=v])}. c .{%Z s. Q Z (s[Loc Y::=v])} ==> G|-{P}. LOCAL Y:=a IN c .{Q}" apply (rule export_s) apply (rule hoare_derivs.Local [THEN conseq1]) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IMPP/Natural.thy --- a/src/HOL/IMPP/Natural.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IMPP/Natural.thy Thu Feb 15 13:04:36 2018 +0100 @@ -103,7 +103,7 @@ declare evaln_elim_cases [elim!] (* evaluation of com is deterministic *) -lemma com_det [rule_format (no_asm)]: " -c-> t ==> (!u. -c-> u --> u=t)" +lemma com_det [rule_format (no_asm)]: " -c-> t \ (\u. -c-> u \ u=t)" apply (erule evalc.induct) apply (erule_tac [8] V = " -c-> s2" for c in thin_rl) apply (blast elim: evalc_WHILE_case)+ @@ -121,7 +121,7 @@ apply blast done -lemma evaln_nonstrict [rule_format]: " -n-> t ==> !m. n<=m --> -m-> t" +lemma evaln_nonstrict [rule_format]: " -n-> t \ \m. n<=m \ -m-> t" apply (erule evaln.induct) apply (auto elim!: Suc_le_D_lemma) done @@ -132,12 +132,12 @@ done lemma evaln_max2: "[| -n1-> t1; -n2-> t2 |] ==> - ? n. -n -> t1 & -n -> t2" + \n. -n -> t1 \ -n -> t2" apply (cut_tac m = "n1" and n = "n2" in nat_le_linear) apply (blast dest: evaln_nonstrict) done -lemma evalc_evaln: " -c-> t ==> ? n. -n-> t" +lemma evalc_evaln: " -c-> t \ \n. -n-> t" apply (erule evalc.induct) apply (tactic \ALLGOALS (REPEAT o eresolve_tac @{context} [exE])\) apply (tactic \TRYALL (EVERY' [dresolve_tac @{context} @{thms evaln_max2}, assume_tac @{context}, @@ -147,7 +147,7 @@ resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW assume_tac @{context})\) done -lemma eval_eq: " -c-> t = (? n. -n-> t)" +lemma eval_eq: " -c-> t = (\n. -n-> t)" apply (fast elim: evalc_evaln evaln_evalc) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IOA/Asig.thy --- a/src/HOL/IOA/Asig.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IOA/Asig.thy Thu Feb 15 13:04:36 2018 +0100 @@ -9,40 +9,40 @@ imports Main begin -type_synonym 'a signature = "('a set * 'a set * 'a set)" +type_synonym 'a signature = "('a set \ 'a set \ 'a set)" -definition "inputs" :: "'action signature => 'action set" - where asig_inputs_def: "inputs == fst" +definition "inputs" :: "'action signature \ 'action set" + where asig_inputs_def: "inputs \ fst" -definition "outputs" :: "'action signature => 'action set" - where asig_outputs_def: "outputs == (fst o snd)" +definition "outputs" :: "'action signature \ 'action set" + where asig_outputs_def: "outputs \ (fst \ snd)" -definition "internals" :: "'action signature => 'action set" - where asig_internals_def: "internals == (snd o snd)" +definition "internals" :: "'action signature \ 'action set" + where asig_internals_def: "internals \ (snd \ snd)" -definition "actions" :: "'action signature => 'action set" - where actions_def: "actions(asig) == (inputs(asig) Un outputs(asig) Un internals(asig))" +definition "actions" :: "'action signature \ 'action set" + where actions_def: "actions(asig) \ (inputs(asig) \ outputs(asig) \ internals(asig))" -definition externals :: "'action signature => 'action set" - where externals_def: "externals(asig) == (inputs(asig) Un outputs(asig))" +definition externals :: "'action signature \ 'action set" + where externals_def: "externals(asig) \ (inputs(asig) \ outputs(asig))" definition is_asig :: "'action signature => bool" - where "is_asig(triple) == - ((inputs(triple) Int outputs(triple) = {}) & - (outputs(triple) Int internals(triple) = {}) & - (inputs(triple) Int internals(triple) = {}))" + where "is_asig(triple) \ + ((inputs(triple) \ outputs(triple) = {}) \ + (outputs(triple) \ internals(triple) = {}) \ + (inputs(triple) \ internals(triple) = {}))" -definition mk_ext_asig :: "'action signature => 'action signature" - where "mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})" +definition mk_ext_asig :: "'action signature \ 'action signature" + where "mk_ext_asig(triple) \ (inputs(triple), outputs(triple), {})" lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def -lemma int_and_ext_is_act: "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)" +lemma int_and_ext_is_act: "\a\internals(S); a\externals(S)\ \ a\actions(S)" apply (simp add: externals_def actions_def) done -lemma ext_is_act: "[|a:externals(S)|] ==> a:actions(S)" +lemma ext_is_act: "a\externals(S) \ a\actions(S)" apply (simp add: externals_def actions_def) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IOA/IOA.thy --- a/src/HOL/IOA/IOA.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IOA/IOA.thy Thu Feb 15 13:04:36 2018 +0100 @@ -18,9 +18,9 @@ (* IO automata *) definition state_trans :: "['action signature, ('action,'state)transition set] => bool" - where "state_trans asig R == - (!triple. triple:R --> fst(snd(triple)):actions(asig)) & - (!a. (a:inputs(asig)) --> (!s1. ? s2. (s1,a,s2):R))" + where "state_trans asig R \ + (\triple. triple \ R \ fst(snd(triple)) \ actions(asig)) \ + (\a. (a \ inputs(asig)) \ (\s1. \s2. (s1,a,s2) \ R))" definition asig_of :: "('action,'state)ioa => 'action signature" where "asig_of == fst" @@ -43,60 +43,60 @@ the first is the action options, the second the state sequence. Finite executions have None actions from some point on. *) definition is_execution_fragment :: "[('action,'state)ioa, ('action,'state)execution] => bool" - where "is_execution_fragment A ex == + where "is_execution_fragment A ex \ let act = fst(ex); state = snd(ex) - in !n a. (act(n)=None --> state(Suc(n)) = state(n)) & - (act(n)=Some(a) --> (state(n),a,state(Suc(n))):trans_of(A))" + in \n a. (act(n)=None \ state(Suc(n)) = state(n)) \ + (act(n)=Some(a) \ (state(n),a,state(Suc(n))) \ trans_of(A))" definition executions :: "('action,'state)ioa => ('action,'state)execution set" - where "executions(ioa) == {e. snd e 0:starts_of(ioa) & is_execution_fragment ioa e}" + where "executions(ioa) \ {e. snd e 0 \ starts_of(ioa) \ is_execution_fragment ioa e}" definition reachable :: "[('action,'state)ioa, 'state] => bool" - where "reachable ioa s == (? ex:executions(ioa). ? n. (snd ex n) = s)" + where "reachable ioa s \ (\ex\executions(ioa). \n. (snd ex n) = s)" definition invariant :: "[('action,'state)ioa, 'state=>bool] => bool" - where "invariant A P == (!s. reachable A s --> P(s))" + where "invariant A P \ (\s. reachable A s \ P(s))" (* Composition of action signatures and automata *) consts - compatible_asigs ::"('a => 'action signature) => bool" - asig_composition ::"('a => 'action signature) => 'action signature" - compatible_ioas ::"('a => ('action,'state)ioa) => bool" - ioa_composition ::"('a => ('action, 'state)ioa) =>('action,'a => 'state)ioa" + compatible_asigs ::"('a \ 'action signature) \ bool" + asig_composition ::"('a \ 'action signature) \ 'action signature" + compatible_ioas ::"('a \ ('action,'state)ioa) \ bool" + ioa_composition ::"('a \ ('action, 'state)ioa) \ ('action,'a \ 'state)ioa" (* binary composition of action signatures and automata *) definition compat_asigs ::"['action signature, 'action signature] => bool" where "compat_asigs a1 a2 == - (((outputs(a1) Int outputs(a2)) = {}) & - ((internals(a1) Int actions(a2)) = {}) & + (((outputs(a1) Int outputs(a2)) = {}) \ + ((internals(a1) Int actions(a2)) = {}) \ ((internals(a2) Int actions(a1)) = {}))" -definition compat_ioas ::"[('action,'s)ioa, ('action,'t)ioa] => bool" - where "compat_ioas ioa1 ioa2 == compat_asigs (asig_of(ioa1)) (asig_of(ioa2))" +definition compat_ioas ::"[('action,'s)ioa, ('action,'t)ioa] \ bool" + where "compat_ioas ioa1 ioa2 \ compat_asigs (asig_of(ioa1)) (asig_of(ioa2))" -definition asig_comp :: "['action signature, 'action signature] => 'action signature" - where "asig_comp a1 a2 == - (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)), - (outputs(a1) Un outputs(a2)), - (internals(a1) Un internals(a2))))" +definition asig_comp :: "['action signature, 'action signature] \ 'action signature" + where "asig_comp a1 a2 \ + (((inputs(a1) \ inputs(a2)) - (outputs(a1) \ outputs(a2)), + (outputs(a1) \ outputs(a2)), + (internals(a1) \ internals(a2))))" -definition par :: "[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa" (infixr "||" 10) - where "(ioa1 || ioa2) == +definition par :: "[('a,'s)ioa, ('a,'t)ioa] \ ('a,'s*'t)ioa" (infixr "||" 10) + where "(ioa1 || ioa2) \ (asig_comp (asig_of ioa1) (asig_of ioa2), - {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)}, + {pr. fst(pr) \ starts_of(ioa1) \ snd(pr) \ starts_of(ioa2)}, {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) - in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) & - (if a:actions(asig_of(ioa1)) then - (fst(s),a,fst(t)):trans_of(ioa1) + in (a \ actions(asig_of(ioa1)) | a \ actions(asig_of(ioa2))) & + (if a \ actions(asig_of(ioa1)) then + (fst(s),a,fst(t)) \ trans_of(ioa1) else fst(t) = fst(s)) & - (if a:actions(asig_of(ioa2)) then - (snd(s),a,snd(t)):trans_of(ioa2) + (if a \ actions(asig_of(ioa2)) then + (snd(s),a,snd(t)) \ trans_of(ioa2) else snd(t) = snd(s))})" @@ -104,35 +104,35 @@ (* Restrict the trace to those members of the set s *) definition filter_oseq :: "('a => bool) => 'a oseq => 'a oseq" - where "filter_oseq p s == - (%i. case s(i) - of None => None - | Some(x) => if p x then Some x else None)" + where "filter_oseq p s \ + (\i. case s(i) + of None \ None + | Some(x) \ if p x then Some x else None)" -definition mk_trace :: "[('action,'state)ioa, 'action oseq] => 'action oseq" - where "mk_trace(ioa) == filter_oseq(%a. a:externals(asig_of(ioa)))" +definition mk_trace :: "[('action,'state)ioa, 'action oseq] \ 'action oseq" + where "mk_trace(ioa) \ filter_oseq(\a. a \ externals(asig_of(ioa)))" (* Does an ioa have an execution with the given trace *) -definition has_trace :: "[('action,'state)ioa, 'action oseq] => bool" - where "has_trace ioa b == (? ex:executions(ioa). b = mk_trace ioa (fst ex))" +definition has_trace :: "[('action,'state)ioa, 'action oseq] \ bool" + where "has_trace ioa b \ (\ex\executions(ioa). b = mk_trace ioa (fst ex))" definition NF :: "'a oseq => 'a oseq" - where "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) & - (!j. j ~: range(f) --> nf(j)= None) & - (!i. nf(i)=None --> (nf (Suc i)) = None)" + where "NF(tr) \ SOME nf. \f. mono(f) \ (\i. nf(i)=tr(f(i))) \ + (\j. j \ range(f) \ nf(j)= None) & + (\i. nf(i)=None --> (nf (Suc i)) = None)" (* All the traces of an ioa *) definition traces :: "('action,'state)ioa => 'action oseq set" - where "traces(ioa) == {trace. ? tr. trace=NF(tr) & has_trace ioa tr}" + where "traces(ioa) \ {trace. \tr. trace=NF(tr) \ has_trace ioa tr}" definition restrict_asig :: "['a signature, 'a set] => 'a signature" - where "restrict_asig asig actns == - (inputs(asig) Int actns, outputs(asig) Int actns, - internals(asig) Un (externals(asig) - actns))" + where "restrict_asig asig actns \ + (inputs(asig) \ actns, outputs(asig) \ actns, + internals(asig) \ (externals(asig) - actns))" definition restrict :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa" - where "restrict ioa actns == + where "restrict ioa actns \ (restrict_asig (asig_of ioa) actns, starts_of(ioa), trans_of(ioa))" @@ -140,23 +140,23 @@ (* Notions of correctness *) definition ioa_implements :: "[('action,'state1)ioa, ('action,'state2)ioa] => bool" - where "ioa_implements ioa1 ioa2 == - ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) & - (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) & - traces(ioa1) <= traces(ioa2))" + where "ioa_implements ioa1 ioa2 \ + ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) \ + (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) \ + traces(ioa1) \ traces(ioa2))" (* Instantiation of abstract IOA by concrete actions *) -definition rename :: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa" - where "rename ioa ren == - (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))}, - {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))}, - {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}), +definition rename :: "('a, 'b)ioa \ ('c \ 'a option) \ ('c,'b)ioa" + where "rename ioa ren \ + (({b. \x. Some(x)= ren(b) \ x \ inputs(asig_of(ioa))}, + {b. \x. Some(x)= ren(b) \ x \ outputs(asig_of(ioa))}, + {b. \x. Some(x)= ren(b) \ x \ internals(asig_of(ioa))}), starts_of(ioa) , {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr)) in - ? x. Some(x) = ren(a) & (s,x,t):trans_of(ioa)})" + \x. Some(x) = ren(a) \ (s,x,t) \ trans_of(ioa)})" declare Let_def [simp] @@ -170,7 +170,7 @@ done lemma trans_in_actions: - "[| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))" + "[| IOA(A); (s1,a,s2) \ trans_of(A) |] ==> a \ actions(asig_of(A))" apply (unfold IOA_def state_trans_def actions_def is_asig_def) apply (erule conjE)+ apply (erule allE, erule impE, assumption) @@ -187,16 +187,16 @@ lemma mk_trace_thm: "(mk_trace A s n = None) = - (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) + (s(n)=None | (\a. s(n)=Some(a) \ a \ externals(asig_of(A)))) & (mk_trace A s n = Some(a)) = - (s(n)=Some(a) & a : externals(asig_of(A)))" + (s(n)=Some(a) \ a \ externals(asig_of(A)))" apply (unfold mk_trace_def filter_oseq_def) apply (case_tac "s n") apply auto done -lemma reachable_0: "s:starts_of(A) ==> reachable A s" +lemma reachable_0: "s \ starts_of(A) \ reachable A s" apply (unfold reachable_def) apply (rule_tac x = "(%i. None, %i. s)" in bexI) apply simp @@ -204,7 +204,7 @@ done lemma reachable_n: - "!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t" + "\A. [| reachable A s; (s,a,t) \ trans_of(A) |] ==> reachable A t" apply (unfold reachable_def exec_rws) apply (simp del: bex_simps) apply (simp (no_asm_simp) only: split_tupled_all) @@ -219,8 +219,8 @@ lemma invariantI: - assumes p1: "!!s. s:starts_of(A) ==> P(s)" - and p2: "!!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t)" + assumes p1: "\s. s \ starts_of(A) \ P(s)" + and p2: "\s t a. [|reachable A s; P(s)|] ==> (s,a,t) \ trans_of(A) \ P(t)" shows "invariant A P" apply (unfold invariant_def reachable_def Let_def exec_rws) apply safe @@ -237,8 +237,8 @@ done lemma invariantI1: - "[| !!s. s : starts_of(A) ==> P(s); - !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) + "[| \s. s \ starts_of(A) \ P(s); + \s t a. reachable A s \ P(s) \ (s,a,t) \ trans_of(A) \ P(t) |] ==> invariant A P" apply (blast intro!: invariantI) done @@ -250,36 +250,36 @@ done lemma actions_asig_comp: - "actions(asig_comp a b) = actions(a) Un actions(b)" + "actions(asig_comp a b) = actions(a) \ actions(b)" apply (auto simp add: actions_def asig_comp_def asig_projections) done lemma starts_of_par: - "starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}" + "starts_of(A || B) = {p. fst(p) \ starts_of(A) \ snd(p) \ starts_of(B)}" apply (simp add: par_def ioa_projections) done (* Every state in an execution is reachable *) lemma states_of_exec_reachable: - "ex:executions(A) ==> !n. reachable A (snd ex n)" + "ex \ executions(A) \ \n. reachable A (snd ex n)" apply (unfold reachable_def) apply fast done lemma trans_of_par4: -"(s,a,t) : trans_of(A || B || C || D) = - ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | - a:actions(asig_of(D))) & - (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) - else fst t=fst s) & - (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) - else fst(snd(t))=fst(snd(s))) & - (if a:actions(asig_of(C)) then - (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) - else fst(snd(snd(t)))=fst(snd(snd(s)))) & - (if a:actions(asig_of(D)) then - (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) +"(s,a,t) \ trans_of(A || B || C || D) = + ((a \ actions(asig_of(A)) | a \ actions(asig_of(B)) | a \ actions(asig_of(C)) | + a \ actions(asig_of(D))) \ + (if a \ actions(asig_of(A)) then (fst(s),a,fst(t)) \ trans_of(A) + else fst t=fst s) \ + (if a \ actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))) \ trans_of(B) + else fst(snd(t))=fst(snd(s))) \ + (if a \ actions(asig_of(C)) then + (fst(snd(snd(s))),a,fst(snd(snd(t)))) \ trans_of(C) + else fst(snd(snd(t)))=fst(snd(snd(s)))) \ + (if a \ actions(asig_of(D)) then + (snd(snd(snd(s))),a,snd(snd(snd(t)))) \ trans_of(D) else snd(snd(snd(t)))=snd(snd(snd(s)))))" (*SLOW*) apply (simp (no_asm) add: par_def actions_asig_comp prod_eq_iff ioa_projections) @@ -298,20 +298,20 @@ lemma externals_of_par: "externals(asig_of(A1||A2)) = - (externals(asig_of(A1)) Un externals(asig_of(A2)))" + (externals(asig_of(A1)) \ externals(asig_of(A2)))" apply (simp add: externals_def asig_of_par asig_comp_def asig_inputs_def asig_outputs_def Un_def set_diff_eq) apply blast done lemma ext1_is_not_int2: - "[| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))" + "[| compat_ioas A1 A2; a \ externals(asig_of(A1))|] ==> a \ internals(asig_of(A2))" apply (unfold externals_def actions_def compat_ioas_def compat_asigs_def) apply auto done lemma ext2_is_not_int1: - "[| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))" + "[| compat_ioas A2 A1 ; a \ externals(asig_of(A1))|] ==> a \ internals(asig_of(A2))" apply (unfold externals_def actions_def compat_ioas_def compat_asigs_def) apply auto done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/IOA/Solve.thy --- a/src/HOL/IOA/Solve.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/IOA/Solve.thy Thu Feb 15 13:04:36 2018 +0100 @@ -9,20 +9,20 @@ imports IOA begin -definition is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool" where - "is_weak_pmap f C A == - (!s:starts_of(C). f(s):starts_of(A)) & - (!s t a. reachable C s & - (s,a,t):trans_of(C) - --> (if a:externals(asig_of(C)) then - (f(s),a,f(t)):trans_of(A) +definition is_weak_pmap :: "['c \ 'a, ('action,'c)ioa,('action,'a)ioa] \ bool" where + "is_weak_pmap f C A \ + (\s\starts_of(C). f(s)\starts_of(A)) \ + (\s t a. reachable C s \ + (s,a,t)\trans_of(C) + \ (if a\externals(asig_of(C)) then + (f(s),a,f(t))\trans_of(A) else f(s)=f(t)))" declare mk_trace_thm [simp] trans_in_actions [simp] lemma trace_inclusion: "[| IOA(C); IOA(A); externals(asig_of(C)) = externals(asig_of(A)); - is_weak_pmap f C A |] ==> traces(C) <= traces(A)" + is_weak_pmap f C A |] ==> traces(C) \ traces(A)" apply (unfold is_weak_pmap_def traces_def) apply (simp (no_asm) add: has_trace_def) @@ -34,7 +34,7 @@ apply simp (* give execution of abstract automata *) - apply (rule_tac x = "(mk_trace A ex1,%i. f (ex2 i))" in bexI) + apply (rule_tac x = "(mk_trace A ex1,\i. f (ex2 i))" in bexI) (* Traces coincide *) apply (simp (no_asm_simp) add: mk_trace_def filter_oseq_idemp) @@ -62,16 +62,16 @@ (* Lemmata *) -lemma imp_conj_lemma: "(P ==> Q-->R) ==> P&Q --> R" +lemma imp_conj_lemma: "(P \ Q\R) \ P\Q \ R" by blast (* fist_order_tautology of externals_of_par *) lemma externals_of_par_extra: - "a:externals(asig_of(A1||A2)) = - (a:externals(asig_of(A1)) & a:externals(asig_of(A2)) | - a:externals(asig_of(A1)) & a~:externals(asig_of(A2)) | - a~:externals(asig_of(A1)) & a:externals(asig_of(A2)))" + "a\externals(asig_of(A1||A2)) = + (a\externals(asig_of(A1)) \ a\externals(asig_of(A2)) \ + a\externals(asig_of(A1)) \ a\externals(asig_of(A2)) \ + a\externals(asig_of(A1)) \ a\externals(asig_of(A2)))" apply (auto simp add: externals_def asig_of_par asig_comp_def asig_inputs_def asig_outputs_def) done @@ -79,7 +79,7 @@ apply (simp add: reachable_def) apply (erule bexE) apply (rule_tac x = - "(filter_oseq (%a. a:actions (asig_of (C1))) (fst ex) , %i. fst (snd ex i))" in bexI) + "(filter_oseq (\a. a\actions (asig_of (C1))) (fst ex) , \i. fst (snd ex i))" in bexI) (* fst(s) is in projected execution *) apply force (* projected execution is indeed an execution *) @@ -96,7 +96,7 @@ apply (simp add: reachable_def) apply (erule bexE) apply (rule_tac x = - "(filter_oseq (%a. a:actions (asig_of (C2))) (fst ex) , %i. snd (snd ex i))" in bexI) + "(filter_oseq (\a. a\actions (asig_of (C2))) (fst ex) , \i. snd (snd ex i))" in bexI) (* fst(s) is in projected execution *) apply force (* projected execution is indeed an execution *) @@ -115,7 +115,7 @@ is_weak_pmap g C2 A2; externals(asig_of(A2))=externals(asig_of(C2)); compat_ioas C1 C2; compat_ioas A1 A2 |] - ==> is_weak_pmap (%p.(f(fst(p)),g(snd(p)))) (C1||C2) (A1||A2)" + ==> is_weak_pmap (\p.(f(fst(p)),g(snd(p)))) (C1||C2) (A1||A2)" apply (unfold is_weak_pmap_def) apply (rule conjI) (* start_states *) @@ -139,7 +139,7 @@ apply (simp add: comp1_reachable comp2_reachable ext_is_act ext1_ext2_is_not_act1) (* case 4 a:~e(A1) | a~:e(A2) *) apply (rule impI) - apply (subgoal_tac "a~:externals (asig_of (A1)) & a~:externals (asig_of (A2))") + apply (subgoal_tac "a\externals (asig_of (A1)) & a\externals (asig_of (A2))") (* delete auxiliary subgoal *) prefer 2 apply force @@ -153,7 +153,7 @@ lemma reachable_rename_ioa: "[| reachable (rename C g) s |] ==> reachable C s" apply (simp add: reachable_def) apply (erule bexE) - apply (rule_tac x = "((%i. case (fst ex i) of None => None | Some (x) => g x) ,snd ex)" in bexI) + apply (rule_tac x = "((\i. case (fst ex i) of None \ None | Some (x) => g x) ,snd ex)" in bexI) apply (simp (no_asm)) (* execution is indeed an execution of C *) apply (simp add: executions_def is_execution_fragment_def par_def diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Induct/SList.thy --- a/src/HOL/Induct/SList.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Induct/SList.thy Thu Feb 15 13:04:36 2018 +0100 @@ -49,8 +49,8 @@ list :: "'a item set => 'a item set" for A :: "'a item set" where - NIL_I: "NIL: list A" - | CONS_I: "[| a: A; M: list A |] ==> CONS a M : list A" + NIL_I: "NIL \ list A" + | CONS_I: "[| a \ A; M \ list A |] ==> CONS a M \ list A" definition "List = list (range Leaf)" @@ -68,7 +68,7 @@ definition List_rec :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b" where - "List_rec M c d = wfrec (pred_sexp^+) + "List_rec M c d = wfrec (pred_sexp\<^sup>+) (%g. List_case c (%x y. d x y (g y))) M" @@ -139,10 +139,10 @@ take_0: "take xs 0 = []" | take_Suc: "take xs (Suc n) = list_case [] (%x l. x # take l n) xs" -lemma ListI: "x : list (range Leaf) ==> x : List" +lemma ListI: "x \ list (range Leaf) \ x \ List" by (simp add: List_def) -lemma ListD: "x : List ==> x : list (range Leaf)" +lemma ListD: "x \ List \ x \ list (range Leaf)" by (simp add: List_def) lemma list_unfold: "list(A) = usum {Numb(0)} (uprod A (list(A)))" @@ -225,10 +225,10 @@ lemmas Cons_inject2 = Cons_Cons_eq [THEN iffD1, THEN conjE] -lemma CONS_D: "CONS M N: list(A) ==> M: A & N: list(A)" +lemma CONS_D: "CONS M N \ list(A) \ M \ A & N \ list(A)" by (induct L == "CONS M N" rule: list.induct) auto -lemma sexp_CONS_D: "CONS M N: sexp ==> M: sexp & N: sexp" +lemma sexp_CONS_D: "CONS M N \ sexp \ M \ sexp \ N \ sexp" apply (simp add: CONS_def In1_def) apply (fast dest!: Scons_D) done @@ -237,14 +237,14 @@ (*Reasoning about constructors and their freeness*) -lemma not_CONS_self: "N: list(A) ==> !M. N ~= CONS M N" +lemma not_CONS_self: "N \ list(A) \ \M. N \ CONS M N" apply (erule list.induct) apply simp_all done -lemma not_Cons_self2: "\x. l ~= x#l" +lemma not_Cons_self2: "\x. l \ x#l" by (induct l rule: list_induct) simp_all -lemma neq_Nil_conv2: "(xs ~= []) = (\y ys. xs = y#ys)" +lemma neq_Nil_conv2: "(xs \ []) = (\y ys. xs = y#ys)" by (induct xs rule: list_induct) auto (** Conversion rules for List_case: case analysis operator **) @@ -262,8 +262,8 @@ hold if pred_sexp^+ were changed to pred_sexp. *) lemma List_rec_unfold_lemma: - "(%M. List_rec M c d) == - wfrec (pred_sexp^+) (%g. List_case c (%x y. d x y (g y)))" + "(\M. List_rec M c d) \ + wfrec (pred_sexp\<^sup>+) (\g. List_case c (\x y. d x y (g y)))" by (simp add: List_rec_def) lemmas List_rec_unfold = @@ -273,16 +273,16 @@ (** pred_sexp lemmas **) lemma pred_sexp_CONS_I1: - "[| M: sexp; N: sexp |] ==> (M, CONS M N) : pred_sexp^+" + "[| M \ sexp; N \ sexp |] ==> (M, CONS M N) \ pred_sexp\<^sup>+" by (simp add: CONS_def In1_def) lemma pred_sexp_CONS_I2: - "[| M: sexp; N: sexp |] ==> (N, CONS M N) : pred_sexp^+" + "[| M \ sexp; N \ sexp |] ==> (N, CONS M N) \ pred_sexp\<^sup>+" by (simp add: CONS_def In1_def) lemma pred_sexp_CONS_D: - "(CONS M1 M2, N) : pred_sexp^+ ==> - (M1,N) : pred_sexp^+ & (M2,N) : pred_sexp^+" + "(CONS M1 M2, N) \ pred_sexp\<^sup>+ \ + (M1,N) \ pred_sexp\<^sup>+ \ (M2,N) \ pred_sexp\<^sup>+" apply (frule pred_sexp_subset_Sigma [THEN trancl_subset_Sigma, THEN subsetD]) apply (blast dest!: sexp_CONS_D intro: pred_sexp_CONS_I1 pred_sexp_CONS_I2 trans_trancl [THEN transD]) @@ -297,7 +297,7 @@ done lemma List_rec_CONS [simp]: - "[| M: sexp; N: sexp |] + "[| M \ sexp; N \ sexp |] ==> List_rec (CONS M N) c h = h M N (List_rec N c h)" apply (rule List_rec_unfold [THEN trans]) apply (simp add: pred_sexp_CONS_I2) @@ -322,11 +322,11 @@ (*Type checking. Useful?*) lemma List_rec_type: - "[| M: list(A); + "[| M \ list(A); A<=sexp; - c: C(NIL); - !!x y r. [| x: A; y: list(A); r: C(y) |] ==> h x y r: C(CONS x y) - |] ==> List_rec M c h : C(M :: 'a item)" + c \ C(NIL); + \x y r. [| x \ A; y \ list(A); r \ C(y) |] ==> h x y r \ C(CONS x y) + |] ==> List_rec M c h \ C(M :: 'a item)" apply (erule list.induct, simp) apply (insert list_subset_sexp) apply (subst List_rec_CONS, blast+) @@ -343,7 +343,7 @@ "Rep_map f(x#xs) = CONS(f x)(Rep_map f xs)" by (simp add: Rep_map_def) -lemma Rep_map_type: "(!!x. f(x): A) ==> Rep_map f xs: list(A)" +lemma Rep_map_type: "(\x. f(x) \ A) \ Rep_map f xs \ list(A)" apply (simp add: Rep_map_def) apply (rule list_induct, auto) done @@ -352,17 +352,17 @@ by (simp add: Abs_map_def) lemma Abs_map_CONS [simp]: - "[| M: sexp; N: sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N" + "[| M \ sexp; N \ sexp |] ==> Abs_map g (CONS M N) = g(M) # Abs_map g N" by (simp add: Abs_map_def) (*Eases the use of primitive recursion.*) lemma def_list_rec_NilCons: - "[| !!xs. f(xs) = list_rec xs c h |] - ==> f [] = c & f(x#xs) = h x xs (f xs)" + "[| \xs. f(xs) = list_rec xs c h |] + ==> f [] = c \ f(x#xs) = h x xs (f xs)" by simp lemma Abs_map_inverse: - "[| M: list(A); A<=sexp; !!z. z: A ==> f(g(z)) = z |] + "[| M \ list(A); A<=sexp; \z. z \ A ==> f(g(z)) = z |] ==> Rep_map f (Abs_map g M) = M" apply (erule list.induct, simp_all) apply (insert list_subset_sexp) @@ -383,7 +383,7 @@ (** list_case **) lemma expand_list_case: - "P(list_case a f xs) = ((xs=[] --> P a ) & (!y ys. xs=y#ys --> P(f y ys)))" + "P(list_case a f xs) = ((xs=[] \ P a ) \ (\y ys. xs=y#ys \ P(f y ys)))" by (induct xs rule: list_induct) simp_all @@ -394,8 +394,8 @@ (** The functional "map" and the generalized functionals **) lemma Abs_Rep_map: - "(!!x. f(x): sexp) ==> - Abs_map g (Rep_map f xs) = map (%t. g(f(t))) xs" + "(\x. f(x)\ sexp) ==> + Abs_map g (Rep_map f xs) = map (\t. g(f(t))) xs" apply (induct xs rule: list_induct) apply (simp_all add: Rep_map_type list_sexp [THEN subsetD]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Induct/Sexp.thy --- a/src/HOL/Induct/Sexp.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Induct/Sexp.thy Thu Feb 15 13:04:36 2018 +0100 @@ -24,9 +24,9 @@ definition sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b, 'a item] => 'b" where - "sexp_case c d e M = (THE z. (EX x. M=Leaf(x) & z=c(x)) - | (EX k. M=Numb(k) & z=d(k)) - | (EX N1 N2. M = Scons N1 N2 & z=e N1 N2))" + "sexp_case c d e M = (THE z. (\x. M=Leaf(x) & z=c(x)) + | (\k. M=Numb(k) & z=d(k)) + | (\N1 N2. M = Scons N1 N2 & z=e N1 N2))" definition pred_sexp :: "('a item * 'a item)set" where diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Isar_Examples/Hoare.thy --- a/src/HOL/Isar_Examples/Hoare.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Isar_Examples/Hoare.thy Thu Feb 15 13:04:36 2018 +0100 @@ -101,7 +101,7 @@ assume cmd: "\ P c Q" fix s s' :: 'a assume sem: "Sem c s s'" - assume "s : P'" with P'P have "s \ P" .. + assume "s \ P'" with P'P have "s \ P" .. with cmd sem have "s' \ Q" .. with QQ' show "s' \ Q'" .. qed @@ -132,7 +132,7 @@ from case_nb show ?thesis proof from sem nb show "Sem c2 s s'" by simp - from s nb show "s : P \ -b" by simp + from s nb show "s \ P \ -b" by simp qed qed qed diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Lattices_Big.thy --- a/src/HOL/Lattices_Big.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Lattices_Big.thy Thu Feb 15 13:04:36 2018 +0100 @@ -711,7 +711,7 @@ let ?A = "insert (Max A) ?B" have "finite ?B" using \finite A\ by simp assume "A \ {}" - with \finite A\ have "Max A : A" by auto + with \finite A\ have "Max A \ A" by auto then have A: "?A = A" using insert_Diff_single insert_absorb by auto then have "P ?B" using \P {}\ step IH [of ?B] by blast moreover @@ -905,7 +905,7 @@ by(induction rule: finite.induct) (auto intro: order.strict_trans) lemma ex_is_arg_min_if_finite: fixes f :: "'a \ 'b :: order" -shows "\ finite S; S \ {} \ \ \x. is_arg_min f (\x. x : S) x" +shows "\ finite S; S \ {} \ \ \x. is_arg_min f (\x. x \ S) x" unfolding is_arg_min_def using ex_min_if_finite[of "f ` S"] by auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/BNF_Corec.thy --- a/src/HOL/Library/BNF_Corec.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/BNF_Corec.thy Thu Feb 15 13:04:36 2018 +0100 @@ -72,7 +72,7 @@ lemma self_bounded_weaken_right: "(a :: 'a :: semilattice_inf) \ inf b a \ a \ b" by (erule le_infE) -lemma symp_iff: "symp R \ R = R^--1" +lemma symp_iff: "symp R \ R = R\\" by (metis antisym conversep.cases conversep_le_swap predicate2I symp_def) lemma equivp_inf: "\equivp R; equivp S\ \ equivp (inf R S)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/Countable_Set.thy --- a/src/HOL/Library/Countable_Set.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/Countable_Set.thy Thu Feb 15 13:04:36 2018 +0100 @@ -266,7 +266,7 @@ by(induction n)(simp_all add: assms) lemma countable_rtrancl: - "(\Y. countable Y \ countable (X `` Y)) \ countable Y \ countable (X^* `` Y)" + "(\Y. countable Y \ countable (X `` Y)) \ countable Y \ countable (X\<^sup>* `` Y)" unfolding rtrancl_is_UN_relpow UN_Image by (intro countable_UN countableI_type countable_relpow) lemma countable_lists[intro, simp]: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/IArray.thy --- a/src/HOL/Library/IArray.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/IArray.thy Thu Feb 15 13:04:36 2018 +0100 @@ -31,10 +31,10 @@ [simp]: "length as = List.length (IArray.list_of as)" qualified fun all :: "('a \ bool) \ 'a iarray \ bool" where -"all p (IArray as) = (ALL a : set as. p a)" +"all p (IArray as) = (\a \ set as. p a)" qualified fun exists :: "('a \ bool) \ 'a iarray \ bool" where -"exists p (IArray as) = (EX a : set as. p a)" +"exists p (IArray as) = (\a \ set as. p a)" lemma list_of_code [code]: "IArray.list_of as = map (\n. as !! n) [0 ..< IArray.length as]" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/Lub_Glb.thy --- a/src/HOL/Library/Lub_Glb.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/Lub_Glb.thy Thu Feb 15 13:04:36 2018 +0100 @@ -11,24 +11,24 @@ text \Thanks to suggestions by James Margetson\ definition setle :: "'a set \ 'a::ord \ bool" (infixl "*<=" 70) - where "S *<= x = (ALL y: S. y \ x)" + where "S *<= x = (\y\S. y \ x)" definition setge :: "'a::ord \ 'a set \ bool" (infixl "<=*" 70) - where "x <=* S = (ALL y: S. x \ y)" + where "x <=* S = (\y\S. x \ y)" subsection \Rules for the Relations \*<=\ and \<=*\\ -lemma setleI: "ALL y: S. y \ x \ S *<= x" +lemma setleI: "\y\S. y \ x \ S *<= x" by (simp add: setle_def) -lemma setleD: "S *<= x \ y: S \ y \ x" +lemma setleD: "S *<= x \ y\S \ y \ x" by (simp add: setle_def) -lemma setgeI: "ALL y: S. x \ y \ x <=* S" +lemma setgeI: "\y\S. x \ y \ x <=* S" by (simp add: setge_def) -lemma setgeD: "x <=* S \ y: S \ x \ y" +lemma setgeD: "x <=* S \ y\S \ x \ y" by (simp add: setge_def) @@ -36,7 +36,7 @@ where "leastP P x = (P x \ x <=* Collect P)" definition isUb :: "'a set \ 'a set \ 'a::ord \ bool" - where "isUb R S x = (S *<= x \ x: R)" + where "isUb R S x = (S *<= x \ x \ R)" definition isLub :: "'a set \ 'a set \ 'a::ord \ bool" where "isLub R S x = leastP (isUb R S) x" @@ -53,19 +53,19 @@ lemma leastPD2: "leastP P x \ x <=* Collect P" by (simp add: leastP_def) -lemma leastPD3: "leastP P x \ y: Collect P \ x \ y" +lemma leastPD3: "leastP P x \ y \ Collect P \ x \ y" by (blast dest!: leastPD2 setgeD) lemma isLubD1: "isLub R S x \ S *<= x" by (simp add: isLub_def isUb_def leastP_def) -lemma isLubD1a: "isLub R S x \ x: R" +lemma isLubD1a: "isLub R S x \ x \ R" by (simp add: isLub_def isUb_def leastP_def) lemma isLub_isUb: "isLub R S x \ isUb R S x" unfolding isUb_def by (blast dest: isLubD1 isLubD1a) -lemma isLubD2: "isLub R S x \ y : S \ y \ x" +lemma isLubD2: "isLub R S x \ y \ S \ y \ x" by (blast dest!: isLubD1 setleD) lemma isLubD3: "isLub R S x \ leastP (isUb R S) x" @@ -77,16 +77,16 @@ lemma isLubI2: "isUb R S x \ x <=* Collect (isUb R S) \ isLub R S x" by (simp add: isLub_def leastP_def) -lemma isUbD: "isUb R S x \ y : S \ y \ x" +lemma isUbD: "isUb R S x \ y \ S \ y \ x" by (simp add: isUb_def setle_def) lemma isUbD2: "isUb R S x \ S *<= x" by (simp add: isUb_def) -lemma isUbD2a: "isUb R S x \ x: R" +lemma isUbD2a: "isUb R S x \ x \ R" by (simp add: isUb_def) -lemma isUbI: "S *<= x \ x: R \ isUb R S x" +lemma isUbI: "S *<= x \ x \ R \ isUb R S x" by (simp add: isUb_def) lemma isLub_le_isUb: "isLub R S x \ isUb R S y \ x \ y" @@ -109,7 +109,7 @@ where "greatestP P x = (P x \ Collect P *<= x)" definition isLb :: "'a set \ 'a set \ 'a::ord \ bool" - where "isLb R S x = (x <=* S \ x: R)" + where "isLb R S x = (x <=* S \ x \ R)" definition isGlb :: "'a set \ 'a set \ 'a::ord \ bool" where "isGlb R S x = greatestP (isLb R S) x" @@ -126,19 +126,19 @@ lemma greatestPD2: "greatestP P x \ Collect P *<= x" by (simp add: greatestP_def) -lemma greatestPD3: "greatestP P x \ y: Collect P \ x \ y" +lemma greatestPD3: "greatestP P x \ y \ Collect P \ x \ y" by (blast dest!: greatestPD2 setleD) lemma isGlbD1: "isGlb R S x \ x <=* S" by (simp add: isGlb_def isLb_def greatestP_def) -lemma isGlbD1a: "isGlb R S x \ x: R" +lemma isGlbD1a: "isGlb R S x \ x \ R" by (simp add: isGlb_def isLb_def greatestP_def) lemma isGlb_isLb: "isGlb R S x \ isLb R S x" unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a) -lemma isGlbD2: "isGlb R S x \ y : S \ y \ x" +lemma isGlbD2: "isGlb R S x \ y \ S \ y \ x" by (blast dest!: isGlbD1 setgeD) lemma isGlbD3: "isGlb R S x \ greatestP (isLb R S) x" @@ -150,16 +150,16 @@ lemma isGlbI2: "isLb R S x \ Collect (isLb R S) *<= x \ isGlb R S x" by (simp add: isGlb_def greatestP_def) -lemma isLbD: "isLb R S x \ y : S \ y \ x" +lemma isLbD: "isLb R S x \ y \ S \ y \ x" by (simp add: isLb_def setge_def) lemma isLbD2: "isLb R S x \ x <=* S " by (simp add: isLb_def) -lemma isLbD2a: "isLb R S x \ x: R" +lemma isLbD2a: "isLb R S x \ x \ R" by (simp add: isLb_def) -lemma isLbI: "x <=* S \ x: R \ isLb R S x" +lemma isLbI: "x <=* S \ x \ R \ isLb R S x" by (simp add: isLb_def) lemma isGlb_le_isLb: "isGlb R S x \ isLb R S y \ x \ y" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/Old_Datatype.thy --- a/src/HOL/Library/Old_Datatype.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/Old_Datatype.thy Thu Feb 15 13:04:36 2018 +0100 @@ -59,7 +59,7 @@ definition ndepth :: "('a, 'b) node => nat" where "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)" definition ntrunc :: "[nat, ('a, 'b) dtree] => ('a, 'b) dtree" - where "ntrunc k N == {n. n:N \ ndepth(n)N \ ndepth(n) ('a, 'b) dtree set" @@ -115,10 +115,10 @@ (*** Introduction rules for Node ***) -lemma Node_K0_I: "(%k. Inr 0, a) : Node" +lemma Node_K0_I: "(\k. Inr 0, a) \ Node" by (simp add: Node_def) -lemma Node_Push_I: "p: Node ==> apfst (Push i) p : Node" +lemma Node_Push_I: "p \ Node \ apfst (Push i) p \ Node" apply (simp add: Node_def Push_def) apply (fast intro!: apfst_conv nat.case(2)[THEN trans]) done @@ -299,35 +299,35 @@ (*** Cartesian Product ***) -lemma uprodI [intro!]: "[| M:A; N:B |] ==> Scons M N : uprod A B" +lemma uprodI [intro!]: "\M\A; N\B\ \ Scons M N \ uprod A B" by (simp add: uprod_def) (*The general elimination rule*) lemma uprodE [elim!]: - "[| c : uprod A B; - !!x y. [| x:A; y:B; c = Scons x y |] ==> P - |] ==> P" + "\c \ uprod A B; + \x y. \x \ A; y \ B; c = Scons x y\ \ P + \ \ P" by (auto simp add: uprod_def) (*Elimination of a pair -- introduces no eigenvariables*) -lemma uprodE2: "[| Scons M N : uprod A B; [| M:A; N:B |] ==> P |] ==> P" +lemma uprodE2: "\Scons M N \ uprod A B; \M \ A; N \ B\ \ P\ \ P" by (auto simp add: uprod_def) (*** Disjoint Sum ***) -lemma usum_In0I [intro]: "M:A ==> In0(M) : usum A B" +lemma usum_In0I [intro]: "M \ A \ In0(M) \ usum A B" by (simp add: usum_def) -lemma usum_In1I [intro]: "N:B ==> In1(N) : usum A B" +lemma usum_In1I [intro]: "N \ B \ In1(N) \ usum A B" by (simp add: usum_def) lemma usumE [elim!]: - "[| u : usum A B; - !!x. [| x:A; u=In0(x) |] ==> P; - !!y. [| y:B; u=In1(y) |] ==> P - |] ==> P" + "\u \ usum A B; + \x. \x \ A; u=In0(x)\ \ P; + \y. \y \ B; u=In1(y)\ \ P + \ \ P" by (auto simp add: usum_def) @@ -440,31 +440,31 @@ (*** Equality for Cartesian Product ***) lemma dprodI [intro!]: - "[| (M,M'):r; (N,N'):s |] ==> (Scons M N, Scons M' N') : dprod r s" + "\(M,M') \ r; (N,N') \ s\ \ (Scons M N, Scons M' N') \ dprod r s" by (auto simp add: dprod_def) (*The general elimination rule*) lemma dprodE [elim!]: - "[| c : dprod r s; - !!x y x' y'. [| (x,x') : r; (y,y') : s; - c = (Scons x y, Scons x' y') |] ==> P - |] ==> P" + "\c \ dprod r s; + \x y x' y'. \(x,x') \ r; (y,y') \ s; + c = (Scons x y, Scons x' y')\ \ P + \ \ P" by (auto simp add: dprod_def) (*** Equality for Disjoint Sum ***) -lemma dsum_In0I [intro]: "(M,M'):r ==> (In0(M), In0(M')) : dsum r s" +lemma dsum_In0I [intro]: "(M,M') \ r \ (In0(M), In0(M')) \ dsum r s" by (auto simp add: dsum_def) -lemma dsum_In1I [intro]: "(N,N'):s ==> (In1(N), In1(N')) : dsum r s" +lemma dsum_In1I [intro]: "(N,N') \ s \ (In1(N), In1(N')) \ dsum r s" by (auto simp add: dsum_def) lemma dsumE [elim!]: - "[| w : dsum r s; - !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P; - !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P - |] ==> P" + "\w \ dsum r s; + \x x'. \ (x,x') \ r; w = (In0(x), In0(x')) \ \ P; + \y y'. \ (y,y') \ s; w = (In1(y), In1(y')) \ \ P + \ \ P" by (auto simp add: dsum_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/Stream.thy --- a/src/HOL/Library/Stream.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/Stream.thy Thu Feb 15 13:04:36 2018 +0100 @@ -457,7 +457,7 @@ lemma sset_flat[simp]: "\xs \ sset s. xs \ [] \ sset (flat s) = (\xs \ sset s. set xs)" (is "?P \ ?L = ?R") proof safe - fix x assume ?P "x : ?L" + fix x assume ?P "x \ ?L" then obtain m where "x = flat s !! m" by (metis image_iff sset_range) with \?P\ obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)" proof (atomize_elim, induct m arbitrary: s rule: less_induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/Sublist.thy --- a/src/HOL/Library/Sublist.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/Sublist.thy Thu Feb 15 13:04:36 2018 +0100 @@ -298,7 +298,7 @@ (is "_ \ \ps. ?P L ps") proof(induction "LEAST n. \xs \L. n = length xs" arbitrary: L) case 0 - have "[] : L" using "0.hyps" LeastI[of "\n. \xs\L. n = length xs"] \L \ {}\ + have "[] \ L" using "0.hyps" LeastI[of "\n. \xs\L. n = length xs"] \L \ {}\ by auto hence "?P L []" by(auto) thus ?case .. diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Library/While_Combinator.thy --- a/src/HOL/Library/While_Combinator.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Library/While_Combinator.thy Thu Feb 15 13:04:36 2018 +0100 @@ -387,7 +387,7 @@ qualified fun rtrancl_while_invariant :: "'a list \ 'a set \ bool" where "rtrancl_while_invariant (ws, Z) = (x \ Z \ set ws \ Z \ distinct ws \ {(x,y). y \ set(f x)} `` (Z - set ws) \ Z \ - Z \ {(x,y). y \ set(f x)}^* `` {x} \ (\z\Z - set ws. p z))" + Z \ {(x,y). y \ set(f x)}\<^sup>* `` {x} \ (\z\Z - set ws. p z))" qualified lemma rtrancl_while_invariant: assumes inv: "rtrancl_while_invariant st" and test: "rtrancl_while_test st" @@ -400,8 +400,8 @@ lemma rtrancl_while_Some: assumes "rtrancl_while = Some(ws,Z)" shows "if ws = [] - then Z = {(x,y). y \ set(f x)}^* `` {x} \ (\z\Z. p z) - else \p(hd ws) \ hd ws \ {(x,y). y \ set(f x)}^* `` {x}" + then Z = {(x,y). y \ set(f x)}\<^sup>* `` {x} \ (\z\Z. p z) + else \p(hd ws) \ hd ws \ {(x,y). y \ set(f x)}\<^sup>* `` {x}" proof - have "rtrancl_while_invariant ([x],{x})" by simp with rtrancl_while_invariant have I: "rtrancl_while_invariant (ws,Z)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/List.thy --- a/src/HOL/List.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/List.thy Thu Feb 15 13:04:36 2018 +0100 @@ -834,7 +834,7 @@ lemma length_greater_0_conv [iff]: "(0 < length xs) = (xs \ [])" by (induct xs) auto -lemma length_pos_if_in_set: "x : set xs \ length xs > 0" +lemma length_pos_if_in_set: "x \ set xs \ length xs > 0" by auto lemma length_Suc_conv: @@ -1118,7 +1118,7 @@ lemma rev_map: "rev (map f xs) = map f (rev xs)" by (induct xs) auto -lemma map_eq_conv[simp]: "(map f xs = map g xs) = (!x : set xs. f x = g x)" +lemma map_eq_conv[simp]: "(map f xs = map g xs) = (\x \ set xs. f x = g x)" by (induct xs) auto lemma map_cong [fundef_cong]: @@ -1296,7 +1296,7 @@ lemma set_append [simp]: "set (xs @ ys) = (set xs \ set ys)" by (induct xs) auto -lemma hd_in_set[simp]: "xs \ [] \ hd xs : set xs" +lemma hd_in_set[simp]: "xs \ [] \ hd xs \ set xs" by(cases xs) auto lemma set_subset_Cons: "set xs \ set (x # xs)" @@ -1317,14 +1317,14 @@ lemma set_map [simp]: "set (map f xs) = f`(set xs)" by (induct xs) auto -lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \ P x}" +lemma set_filter [simp]: "set (filter P xs) = {x. x \ set xs \ P x}" by (induct xs) auto lemma set_upt [simp]: "set[i.. \ys zs. xs = ys @ x # zs" +lemma split_list: "x \ set xs \ \ys zs. xs = ys @ x # zs" proof (induct xs) case Nil thus ?case by simp next @@ -1334,7 +1334,7 @@ lemma in_set_conv_decomp: "x \ set xs \ (\ys zs. xs = ys @ x # zs)" by (auto elim: split_list) -lemma split_list_first: "x : set xs \ \ys zs. xs = ys @ x # zs \ x \ set ys" +lemma split_list_first: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set ys" proof (induct xs) case Nil thus ?case by simp next @@ -1348,7 +1348,7 @@ qed lemma in_set_conv_decomp_first: - "(x : set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)" + "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set ys)" by (auto dest!: split_list_first) lemma split_list_last: "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs" @@ -1365,7 +1365,7 @@ qed lemma in_set_conv_decomp_last: - "(x : set xs) = (\ys zs. xs = ys @ x # zs \ x \ set zs)" + "(x \ set xs) = (\ys zs. xs = ys @ x # zs \ x \ set zs)" by (auto dest!: split_list_last) lemma split_list_prop: "\x \ set xs. P x \ \ys x zs. xs = ys @ x # zs \ P x" @@ -1761,14 +1761,14 @@ assume "?rhs" then show "?lhs" by simp qed -lemma list_ball_nth: "[| n < length xs; !x : set xs. P x|] ==> P(xs!n)" +lemma list_ball_nth: "\n < length xs; \x \ set xs. P x\ \ P(xs!n)" by (auto simp add: set_conv_nth) -lemma nth_mem [simp]: "n < length xs ==> xs!n : set xs" +lemma nth_mem [simp]: "n < length xs \ xs!n \ set xs" by (auto simp add: set_conv_nth) lemma all_nth_imp_all_set: - "[| !i < length xs. P(xs!i); x : set xs|] ==> P x" + "\!i < length xs. P(xs!i); x \ set xs\ \ P x" by (auto simp add: set_conv_nth) lemma all_set_conv_all_nth: @@ -1873,7 +1873,7 @@ lemma set_update_subset_insert: "set(xs[i:=x]) <= insert x (set xs)" by (induct xs arbitrary: i) (auto split: nat.split) -lemma set_update_subsetI: "[| set xs <= A; x:A |] ==> set(xs[i := x]) <= A" +lemma set_update_subsetI: "\set xs \ A; x \ A\ \ set(xs[i := x]) \ A" by (blast dest!: set_update_subset_insert [THEN subsetD]) lemma set_update_memI: "n < length xs \ x \ set (xs[n := x])" @@ -1948,10 +1948,10 @@ by (induct xs arbitrary: ys) auto lemma append_butlast_last_id [simp]: - "xs \ [] ==> butlast xs @ [last xs] = xs" + "xs \ [] \ butlast xs @ [last xs] = xs" by (induct xs) auto -lemma in_set_butlastD: "x : set (butlast xs) ==> x : set xs" +lemma in_set_butlastD: "x \ set (butlast xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) lemma in_set_butlast_appendI: @@ -2184,10 +2184,10 @@ apply(auto simp:drop_Cons split:nat.split) by (metis set_drop_subset subset_iff) -lemma in_set_takeD: "x : set(take n xs) \ x : set xs" +lemma in_set_takeD: "x \ set(take n xs) \ x \ set xs" using set_take_subset by fast -lemma in_set_dropD: "x : set(drop n xs) \ x : set xs" +lemma in_set_dropD: "x \ set(drop n xs) \ x \ set xs" using set_drop_subset by fast lemma append_eq_conv_conj: @@ -2277,10 +2277,10 @@ by (induct xs) auto lemma takeWhile_append2 [simp]: - "(!!x. x : set xs ==> P x) ==> takeWhile P (xs @ ys) = xs @ takeWhile P ys" + "(\x. x \ set xs \ P x) \ takeWhile P (xs @ ys) = xs @ takeWhile P ys" by (induct xs) auto -lemma takeWhile_tail: "\ P x ==> takeWhile P (xs @ (x#l)) = takeWhile P xs" +lemma takeWhile_tail: "\ P x \ takeWhile P (xs @ (x#l)) = takeWhile P xs" by (induct xs) auto lemma takeWhile_nth: "j < length (takeWhile P xs) \ takeWhile P xs ! j = xs ! j" @@ -2298,7 +2298,7 @@ by (induct xs) auto lemma dropWhile_append2 [simp]: - "(!!x. x:set xs ==> P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys" + "(\x. x \ set xs \ P(x)) ==> dropWhile P (xs @ ys) = dropWhile P ys" by (induct xs) auto lemma dropWhile_append3: @@ -2312,7 +2312,7 @@ lemma set_dropWhileD: "x \ set (dropWhile P xs) \ x \ set xs" by (induct xs) (auto split: if_split_asm) -lemma set_takeWhileD: "x : set (takeWhile P xs) ==> x : set xs \ P x" +lemma set_takeWhileD: "x \ set (takeWhile P xs) \ x \ set xs \ P x" by (induct xs) (auto split: if_split_asm) lemma takeWhile_eq_all_conv[simp]: @@ -2421,13 +2421,13 @@ by(induction xs rule: induct_list012) auto lemma takeWhile_cong [fundef_cong]: - "[| l = k; !!x. x : set l ==> P x = Q x |] - ==> takeWhile P l = takeWhile Q k" + "\l = k; \x. x \ set l \ P x = Q x\ + \ takeWhile P l = takeWhile Q k" by (induct k arbitrary: l) (simp_all) lemma dropWhile_cong [fundef_cong]: - "[| l = k; !!x. x : set l ==> P x = Q x |] - ==> dropWhile P l = dropWhile Q k" + "\l = k; \x. x \ set l \ P x = Q x\ + \ dropWhile P l = dropWhile Q k" by (induct k arbitrary: l, simp_all) lemma takeWhile_idem [simp]: @@ -2590,7 +2590,7 @@ by (induct xs ys rule:list_induct2') auto lemma in_set_zipE: - "(x,y) : set(zip xs ys) \ (\ x : set xs; y : set ys \ \ R) \ R" + "(x,y) \ set(zip xs ys) \ (\ x \ set xs; y \ set ys \ \ R) \ R" by(blast dest: set_zip_leftD set_zip_rightD) lemma zip_map_fst_snd: "zip (map fst zs) (map snd zs) = zs" @@ -3921,19 +3921,19 @@ by (induct zs) auto lemma in_set_remove1[simp]: - "a \ b \ a : set(remove1 b xs) = (a : set xs)" + "a \ b \ a \ set(remove1 b xs) = (a \ set xs)" apply (induct xs) apply auto done -lemma set_remove1_subset: "set(remove1 x xs) <= set xs" +lemma set_remove1_subset: "set(remove1 x xs) \ set xs" apply(induct xs) apply simp apply simp apply blast done -lemma set_remove1_eq [simp]: "distinct xs ==> set(remove1 x xs) = set xs - {x}" +lemma set_remove1_eq [simp]: "distinct xs \ set(remove1 x xs) = set xs - {x}" apply(induct xs) apply simp apply simp @@ -3941,7 +3941,7 @@ done lemma length_remove1: - "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)" + "length(remove1 x xs) = (if x \ set xs then length xs - 1 else length xs)" by (induct xs) (auto dest!:length_pos_if_in_set) lemma remove1_filter_not[simp]: @@ -3952,10 +3952,10 @@ "filter Q (remove1 x xs) = remove1 x (filter Q xs)" by (induct xs) auto -lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)" +lemma notin_set_remove1[simp]: "x \ set xs \ x \ set(remove1 y xs)" by(insert set_remove1_subset) fast -lemma distinct_remove1[simp]: "distinct xs ==> distinct(remove1 x xs)" +lemma distinct_remove1[simp]: "distinct xs \ distinct(remove1 x xs)" by (induct xs) simp_all lemma remove1_remdups: @@ -4000,7 +4000,7 @@ by (simp add: removeAll_filter_not_eq) lemma distinct_remove1_removeAll: - "distinct xs ==> remove1 x xs = removeAll x xs" + "distinct xs \ remove1 x xs = removeAll x xs" by (induct xs) simp_all lemma map_removeAll_inj_on: "inj_on f (insert x (set xs)) \ @@ -4028,7 +4028,8 @@ lemma replicate_eqI: assumes "length xs = n" and "\y. y \ set xs \ y = x" shows "xs = replicate n x" -using assms proof (induct xs arbitrary: n) + using assms +proof (induct xs arbitrary: n) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases n) simp_all @@ -4393,7 +4394,7 @@ by (auto simp add: nths_def) lemma length_nths: - "length (nths xs I) = card{i. i < length xs \ i : I}" + "length (nths xs I) = card{i. i < length xs \ i \ I}" by(simp add: nths_def length_filter_conv_card cong:conj_cong) lemma nths_shift_lemma_Suc: @@ -4407,12 +4408,12 @@ done lemma nths_shift_lemma: - "map fst [p<-zip xs [i.. A] = + map fst [p<-zip xs [0.. A]" by (induct xs rule: rev_induct) (simp_all add: add.commute) lemma nths_append: - "nths (l @ l') A = nths l A @ nths l' {j. j + length l : A}" + "nths (l @ l') A = nths l A @ nths l' {j. j + length l \ A}" apply (unfold nths_def) apply (induct l' rule: rev_induct, simp) apply (simp add: upt_add_eq_append[of 0] nths_shift_lemma) @@ -4420,7 +4421,7 @@ done lemma nths_Cons: - "nths (x # l) A = (if 0:A then [x] else []) @ nths l {j. Suc j : A}" + "nths (x # l) A = (if 0 \ A then [x] else []) @ nths l {j. Suc j \ A}" apply (induct l rule: rev_induct) apply (simp add: nths_def) apply (simp del: append_Cons add: append_Cons[symmetric] nths_append) @@ -4443,7 +4444,7 @@ lemma in_set_nthsD: "x \ set(nths xs I) \ x \ set xs" by(auto simp add:set_nths) -lemma nths_singleton [simp]: "nths [x] A = (if 0 : A then [x] else [])" +lemma nths_singleton [simp]: "nths [x] A = (if 0 \ A then [x] else [])" by (simp add: nths_Cons) lemma distinct_nthsI[simp]: "distinct xs \ distinct (nths xs I)" @@ -4819,7 +4820,7 @@ next case (insertI M xs) then obtain n where "\s\M. length s < n" by blast - hence "ALL s:insert xs M. size s < max n (size xs) + 1" by auto + hence "\s\insert xs M. size s < max n (size xs) + 1" by auto thus ?case .. qed @@ -5654,10 +5655,10 @@ lists :: "'a set => 'a list set" for A :: "'a set" where - Nil [intro!, simp]: "[]: lists A" - | Cons [intro!, simp]: "[| a: A; l: lists A|] ==> a#l : lists A" - -inductive_cases listsE [elim!]: "x#l : lists A" + Nil [intro!, simp]: "[] \ lists A" + | Cons [intro!, simp]: "\a \ A; l \ lists A\ \ a#l \ lists A" + +inductive_cases listsE [elim!]: "x#l \ lists A" inductive_cases listspE [elim!]: "listsp A (x # l)" inductive_simps listsp_simps[code]: @@ -5685,7 +5686,7 @@ lemmas lists_Int_eq [simp] = listsp_inf_eq [to_set] -lemma Cons_in_lists_iff[simp]: "x#xs : lists A \ x:A \ xs : lists A" +lemma Cons_in_lists_iff[simp]: "x#xs \ lists A \ x \ A \ xs \ lists A" by auto lemma append_in_listsp_conv [iff]: @@ -5791,7 +5792,7 @@ done lemma lexn_length: - "(xs, ys) : lexn r n ==> length xs = n \ length ys = n" + "(xs, ys) \ lexn r n \ length xs = n \ length ys = n" by (induct n arbitrary: xs ys) auto lemma wf_lex [intro!]: "wf r ==> wf (lex r)" @@ -5807,7 +5808,7 @@ lemma lexn_conv: "lexn r n = {(xs,ys). length xs = n \ length ys = n \ - (\xys x y xs' ys'. xs= xys @ x#xs' \ ys= xys @ y # ys' \ (x, y):r)}" + (\xys x y xs' ys'. xs= xys @ x#xs' \ ys= xys @ y # ys' \ (x, y) \ r)}" apply (induct n, simp) apply (simp add: image_Collect lex_prod_def, safe, blast) apply (rule_tac x = "ab # xys" in exI, simp) @@ -5883,7 +5884,7 @@ lemma lex_conv: "lex r = {(xs,ys). length xs = length ys \ - (\xys x y xs' ys'. xs = xys @ x # xs' \ ys = xys @ y # ys' \ (x, y):r)}" + (\xys x y xs' ys'. xs = xys @ x # xs' \ ys = xys @ y # ys' \ (x, y) \ r)}" by (force simp add: lex_def lexn_conv) lemma wf_lenlex [intro!]: "wf r ==> wf (lenlex r)" @@ -6364,11 +6365,11 @@ unfolding listrel1_def by blast -lemma listrel1_converse: "listrel1 (r^-1) = (listrel1 r)^-1" +lemma listrel1_converse: "listrel1 (r\) = (listrel1 r)\" unfolding listrel1_def by blast lemma in_listrel1_converse: - "(x,y) : listrel1 (r^-1) \ (x,y) : (listrel1 r)^-1" + "(x,y) \ listrel1 (r\) \ (x,y) \ (listrel1 r)\" unfolding listrel1_def by blast lemma listrel1_iff_update: @@ -6489,13 +6490,13 @@ theorem equiv_listrel: "equiv A r \ equiv (lists A) (listrel r)" by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans) -lemma listrel_rtrancl_refl[iff]: "(xs,xs) : listrel(r^*)" +lemma listrel_rtrancl_refl[iff]: "(xs,xs) \ listrel(r\<^sup>*)" using listrel_refl_on[of UNIV, OF refl_rtrancl] by(auto simp: refl_on_def) lemma listrel_rtrancl_trans: - "\ (xs,ys) : listrel(r^*); (ys,zs) : listrel(r^*) \ - \ (xs,zs) : listrel(r^*)" + "\(xs,ys) \ listrel(r\<^sup>*); (ys,zs) \ listrel(r\<^sup>*)\ + \ (xs,zs) \ listrel(r\<^sup>*)" by (metis listrel_trans trans_def trans_rtrancl) @@ -6509,32 +6510,32 @@ text \Relating @{term listrel1}, @{term listrel} and closures:\ lemma listrel1_rtrancl_subset_rtrancl_listrel1: - "listrel1 (r^*) \ (listrel1 r)^*" + "listrel1 (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof (rule subrelI) - fix xs ys assume 1: "(xs,ys) \ listrel1 (r^*)" + fix xs ys assume 1: "(xs,ys) \ listrel1 (r\<^sup>*)" { fix x y us vs - have "(x,y) : r^* \ (us @ x # vs, us @ y # vs) : (listrel1 r)^*" + have "(x,y) \ r\<^sup>* \ (us @ x # vs, us @ y # vs) \ (listrel1 r)\<^sup>*" proof(induct rule: rtrancl.induct) case rtrancl_refl show ?case by simp next case rtrancl_into_rtrancl thus ?case by (metis listrel1I rtrancl.rtrancl_into_rtrancl) qed } - thus "(xs,ys) \ (listrel1 r)^*" using 1 by(blast elim: listrel1E) + thus "(xs,ys) \ (listrel1 r)\<^sup>*" using 1 by(blast elim: listrel1E) qed -lemma rtrancl_listrel1_eq_len: "(x,y) \ (listrel1 r)^* \ length x = length y" +lemma rtrancl_listrel1_eq_len: "(x,y) \ (listrel1 r)\<^sup>* \ length x = length y" by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len) lemma rtrancl_listrel1_ConsI1: - "(xs,ys) : (listrel1 r)^* \ (x#xs,x#ys) : (listrel1 r)^*" + "(xs,ys) \ (listrel1 r)\<^sup>* \ (x#xs,x#ys) \ (listrel1 r)\<^sup>*" apply(induct rule: rtrancl.induct) apply simp by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl) lemma rtrancl_listrel1_ConsI2: - "(x,y) \ r^* \ (xs, ys) \ (listrel1 r)^* - \ (x # xs, y # ys) \ (listrel1 r)^*" + "(x,y) \ r\<^sup>* \ (xs, ys) \ (listrel1 r)\<^sup>* + \ (x # xs, y # ys) \ (listrel1 r)\<^sup>*" by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1 subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1]) @@ -6543,21 +6544,21 @@ by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def) lemma listrel_reflcl_if_listrel1: - "(xs,ys) : listrel1 r \ (xs,ys) : listrel(r^*)" + "(xs,ys) \ listrel1 r \ (xs,ys) \ listrel(r\<^sup>*)" by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip) -lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r^*) = (listrel1 r)^*" +lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r\<^sup>*) = (listrel1 r)\<^sup>*" proof - { fix x y assume "(x,y) \ listrel (r^*)" - then have "(x,y) \ (listrel1 r)^*" + { fix x y assume "(x,y) \ listrel (r\<^sup>*)" + then have "(x,y) \ (listrel1 r)\<^sup>*" by induct (auto intro: rtrancl_listrel1_ConsI2) } - then show "listrel (r^*) \ (listrel1 r)^*" + then show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" by (rule subrelI) next - show "listrel (r^*) \ (listrel1 r)^*" + show "listrel (r\<^sup>*) \ (listrel1 r)\<^sup>*" proof(rule subrelI) - fix xs ys assume "(xs,ys) \ (listrel1 r)^*" - then show "(xs,ys) \ listrel (r^*)" + fix xs ys assume "(xs,ys) \ (listrel1 r)\<^sup>*" + then show "(xs,ys) \ listrel (r\<^sup>*)" proof induct case base show ?case by(auto simp add: listrel_iff_zip set_zip) next @@ -6568,10 +6569,10 @@ qed lemma rtrancl_listrel1_if_listrel: - "(xs,ys) : listrel r \ (xs,ys) : (listrel1 r)^*" + "(xs,ys) \ listrel r \ (xs,ys) \ (listrel1 r)\<^sup>*" by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI) -lemma listrel_subset_rtrancl_listrel1: "listrel r \ (listrel1 r)^*" +lemma listrel_subset_rtrancl_listrel1: "listrel r \ (listrel1 r)\<^sup>*" by(fast intro:rtrancl_listrel1_if_listrel) @@ -6833,8 +6834,8 @@ show "x \ set xs \ P x" by (metis Cons Cons_eq_filter_iff in_set_conv_decomp set_sort) next - fix y assume "y : set xs \ P y" - hence "y : set (filter P xs)" by auto + fix y assume "y \ set xs \ P y" + hence "y \ set (filter P xs)" by auto thus "x \ y" by (metis Cons eq_iff filter_sort set_ConsD set_sort sorted_Cons sorted_sort) qed @@ -7190,7 +7191,7 @@ by (auto simp add: Id_on_def) lemma [code]: - "R `` S = List.map_project (%(x, y). if x : S then Some y else None) R" + "R `` S = List.map_project (\(x, y). if x \ S then Some y else None) R" unfolding map_project_def by (auto split: prod.split if_split_asm) lemma trancl_set_ntrancl [code]: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Matrix_LP/ComputeFloat.thy --- a/src/HOL/Matrix_LP/ComputeFloat.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Matrix_LP/ComputeFloat.thy Thu Feb 15 13:04:36 2018 +0100 @@ -16,7 +16,7 @@ where "int_of_real x = (SOME y. real_of_int y = x)" definition real_is_int :: "real \ bool" - where "real_is_int x = (EX (u::int). x = real_of_int u)" + where "real_is_int x = (\(u::int). x = real_of_int u)" lemma real_is_int_def2: "real_is_int x = (x = real_of_int (int_of_real x))" by (auto simp add: real_is_int_def int_of_real_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Matrix_LP/Matrix.thy --- a/src/HOL/Matrix_LP/Matrix.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Matrix_LP/Matrix.thy Thu Feb 15 13:04:36 2018 +0100 @@ -138,7 +138,7 @@ thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def) qed -lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \ (Rep_matrix A j i) = 0)" (is "_ = ?st") +lemma ncols_le: "(ncols A <= n) = (\j i. n <= i \ (Rep_matrix A j i) = 0)" (is "_ = ?st") apply (auto) apply (simp add: ncols) proof (simp add: ncols_def, auto) @@ -152,13 +152,13 @@ assume "(a,b) \ ?P" then have "?p \ {}" by (auto) with a have "?m \ ?p" by (simp) - moreover have "!x. (x \ ?p \ (? y. (Rep_matrix A y x) \ 0))" by (simp add: nonzero_positions_def image_def) - ultimately have "? y. (Rep_matrix A y ?m) \ 0" by (simp) + moreover have "\x. (x \ ?p \ (\y. (Rep_matrix A y x) \ 0))" by (simp add: nonzero_positions_def image_def) + ultimately have "\y. (Rep_matrix A y ?m) \ 0" by (simp) moreover assume ?st ultimately show "False" using b by (simp) qed -lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \ 0)" +lemma less_ncols: "(n < ncols A) = (\j i. n <= i & (Rep_matrix A j i) \ 0)" proof - have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith show ?thesis by (simp add: a ncols_le) @@ -172,15 +172,15 @@ apply (drule_tac x="ncols A" in spec) by (simp add: ncols) -lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" (is ?s) +lemma nrows_le: "(nrows A <= n) = (\j i. n <= j \ (Rep_matrix A j i) = 0)" (is ?s) proof - have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp) - also have "\ = (! j i. n <= i \ (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le) - also have "\ = (! j i. n <= i \ (Rep_matrix A i j) = 0)" by (simp) - finally show "(nrows A <= n) = (! j i. n <= j \ (Rep_matrix A j i) = 0)" by (auto) + also have "\ = (\j i. n <= i \ (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le) + also have "\ = (\j i. n <= i \ (Rep_matrix A i j) = 0)" by (simp) + finally show "(nrows A <= n) = (\j i. n <= j \ (Rep_matrix A j i) = 0)" by (auto) qed -lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \ 0)" +lemma less_nrows: "(m < nrows A) = (\j i. m <= j & (Rep_matrix A j i) \ 0)" proof - have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith show ?thesis by (simp add: a nrows_le) @@ -218,13 +218,13 @@ by simp lemma RepAbs_matrix: - assumes aem: "? m. ! j i. m <= j \ x j i = 0" (is ?em) and aen:"? n. ! j i. (n <= i \ x j i = 0)" (is ?en) + assumes aem: "\m. \j i. m <= j \ x j i = 0" (is ?em) and aen:"\n. \j i. (n <= i \ x j i = 0)" (is ?en) shows "(Rep_matrix (Abs_matrix x)) = x" apply (rule Abs_matrix_inverse) apply (simp add: matrix_def nonzero_positions_def) proof - - from aem obtain m where a: "! j i. m <= j \ x j i = 0" by (blast) - from aen obtain n where b: "! j i. n <= i \ x j i = 0" by (blast) + from aem obtain m where a: "\j i. m <= j \ x j i = 0" by (blast) + from aen obtain n where b: "\j i. n <= i \ x j i = 0" by (blast) let ?u = "{(i, j). x i j \ 0}" let ?v = "{(i, j). i < m & j < n}" have c: "!! (m::nat) a. ~(m <= a) \ a < m" by (arith) @@ -260,10 +260,10 @@ by (simp add: combine_infmatrix_def) definition commutative :: "('a \ 'a \ 'b) \ bool" where -"commutative f == ! x y. f x y = f y x" +"commutative f == \x y. f x y = f y x" definition associative :: "('a \ 'a \ 'a) \ bool" where -"associative f == ! x y z. f (f x y) z = f x (f y z)" +"associative f == \x y z. f (f x y) z = f x (f y z)" text\ To reason about associativity and commutativity of operations on matrices, @@ -344,13 +344,13 @@ by (simp add: ncols_le) definition zero_r_neutral :: "('a \ 'b::zero \ 'a) \ bool" where - "zero_r_neutral f == ! a. f a 0 = a" + "zero_r_neutral f == \a. f a 0 = a" definition zero_l_neutral :: "('a::zero \ 'b \ 'b) \ bool" where - "zero_l_neutral f == ! a. f 0 a = a" + "zero_l_neutral f == \a. f 0 a = a" definition zero_closed :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ bool" where - "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)" + "zero_closed f == (\x. f x 0 = 0) & (\y. f 0 y = 0)" primrec foldseq :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" where @@ -365,17 +365,17 @@ lemma foldseq_assoc : "associative f \ foldseq f = foldseq_transposed f" proof - assume a:"associative f" - then have sublemma: "!! n. ! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + then have sublemma: "\n. \N s. N <= n \ foldseq f s N = foldseq_transposed f s N" proof - fix n - show "!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + show "\N s. N <= n \ foldseq f s N = foldseq_transposed f s N" proof (induct n) - show "!N s. N <= 0 \ foldseq f s N = foldseq_transposed f s N" by simp + show "\N s. N <= 0 \ foldseq f s N = foldseq_transposed f s N" by simp next fix n - assume b:"! N s. N <= n \ foldseq f s N = foldseq_transposed f s N" - have c:"!!N s. N <= n \ foldseq f s N = foldseq_transposed f s N" by (simp add: b) - show "! N t. N <= Suc n \ foldseq f t N = foldseq_transposed f t N" + assume b: "\N s. N <= n \ foldseq f s N = foldseq_transposed f s N" + have c:"\N s. N <= n \ foldseq f s N = foldseq_transposed f s N" by (simp add: b) + show "\N t. N <= Suc n \ foldseq f t N = foldseq_transposed f t N" proof (auto) fix N t assume Nsuc: "N <= Suc n" @@ -386,7 +386,7 @@ next assume "~(N <= n)" with Nsuc have Nsuceq: "N = Suc n" by simp - have neqz: "n \ 0 \ ? m. n = Suc m & Suc m <= n" by arith + have neqz: "n \ 0 \ \m. n = Suc m & Suc m <= n" by arith have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def) show "foldseq f t N = foldseq_transposed f t N" apply (simp add: Nsuceq) @@ -426,14 +426,14 @@ from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def) from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def) from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def) - have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" + have "\n. (\u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" apply (induct_tac n) apply (simp+, auto) by (simp add: a b c) then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp qed -theorem "\associative f; associative g; \a b c d. g (f a b) (f c d) = f (g a c) (g b d); ? x y. (f x) \ (f y); ? x y. (g x) \ (g y); f x x = x; g x x = x\ \ f=g | (! y. f y x = y) | (! y. g y x = y)" +theorem "\associative f; associative g; \a b c d. g (f a b) (f c d) = f (g a c) (g b d); \x y. (f x) \ (f y); \x y. (g x) \ (g y); f x x = x; g x x = x\ \ f=g | (\y. f y x = y) | (\y. g y x = y)" oops (* Model found @@ -452,10 +452,10 @@ *) lemma foldseq_zero: -assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \ s i = 0" +assumes fz: "f 0 0 = 0" and sz: "\i. i <= n \ s i = 0" shows "foldseq f s n = 0" proof - - have "!! n. ! s. (! i. i <= n \ s i = 0) \ foldseq f s n = 0" + have "\n. \s. (\i. i <= n \ s i = 0) \ foldseq f s n = 0" apply (induct_tac n) apply (simp) by (simp add: fz) @@ -463,10 +463,10 @@ qed lemma foldseq_significant_positions: - assumes p: "! i. i <= N \ S i = T i" + assumes p: "\i. i <= N \ S i = T i" shows "foldseq f S N = foldseq f T N" proof - - have "!! m . ! s t. (! i. i<=m \ s i = t i) \ foldseq f s m = foldseq f t m" + have "\m. \s t. (\i. i<=m \ s i = t i) \ foldseq f s m = foldseq f t m" apply (induct_tac m) apply (simp) apply (simp) @@ -488,9 +488,9 @@ assumes "M <= N" shows "foldseq f S N = foldseq f (% k. (if k < M then (S k) else (foldseq f (% k. S(k+M)) (N-M)))) M" proof - - have suc: "!! a b. \a <= Suc b; a \ Suc b\ \ a <= b" by arith - have a:"!! a b c . a = b \ f c a = f c b" by blast - have "!! n. ! m s. m <= n \ foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m" + have suc: "\a b. \a <= Suc b; a \ Suc b\ \ a <= b" by arith + have a: "\a b c . a = b \ f c a = f c b" by blast + have "\n. \m s. m <= n \ foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m" apply (induct_tac n) apply (simp) apply (simp) @@ -509,7 +509,7 @@ have subd: "foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m = foldseq f (% k. s(Suc k)) na" by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb) - from subb have sube: "m \ 0 \ ? mm. m = Suc mm & mm <= na" by arith + from subb have sube: "m \ 0 \ \mm. m = Suc mm & mm <= na" by arith show "f (s 0) (foldseq f (\k. if k < m then s (Suc k) else foldseq f (\k. s (Suc (k + m))) (na - m)) m) = foldseq f (\k. if k < m then s k else foldseq f (\k. s (k + m)) (Suc na - m)) m" apply (simp add: subd) @@ -527,7 +527,7 @@ lemma foldseq_zerotail: assumes fz: "f 0 0 = 0" - and sz: "! i. n <= i \ s i = 0" + and sz: "\i. n <= i \ s i = 0" and nm: "n <= m" shows "foldseq f s n = foldseq f s m" @@ -541,15 +541,15 @@ qed lemma foldseq_zerotail2: - assumes "! x. f x 0 = x" - and "! i. n < i \ s i = 0" + assumes "\x. f x 0 = x" + and "\i. n < i \ s i = 0" and nm: "n <= m" shows "foldseq f s n = foldseq f s m" proof - have "f 0 0 = 0" by (simp add: assms) - have b:"!! m n. n <= m \ m \ n \ ? k. m-n = Suc k" by arith + have b: "\m n. n <= m \ m \ n \ \k. m-n = Suc k" by arith have c: "0 <= m" by simp - have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith + have d: "\k. k \ 0 \ \l. k = Suc l" by arith show ?thesis apply (subst foldseq_tail[OF nm]) apply (rule foldseq_significant_positions) @@ -567,10 +567,10 @@ qed lemma foldseq_zerostart: - "! x. f 0 (f 0 x) = f 0 x \ ! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" + "\x. f 0 (f 0 x) = f 0 x \ \i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" proof - - assume f00x: "! x. f 0 (f 0 x) = f 0 x" - have "! s. (! i. i<=n \ s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" + assume f00x: "\x. f 0 (f 0 x) = f 0 x" + have "\s. (\i. i<=n \ s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" apply (induct n) apply (simp) apply (rule allI, rule impI) @@ -578,25 +578,25 @@ fix n fix s have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp - assume b: "! s. ((\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n)))" + assume b: "\s. ((\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n)))" from b have c:"!! s. (\i\n. s i = 0) \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp - assume d: "! i. i <= Suc n \ s i = 0" + assume d: "\i. i <= Suc n \ s i = 0" show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))" apply (subst a) apply (subst c) by (simp add: d f00x)+ qed - then show "! i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp + then show "\i. i <= n \ s i = 0 \ foldseq f s (Suc n) = f 0 (s (Suc n))" by simp qed lemma foldseq_zerostart2: - "! x. f 0 x = x \ ! i. i < n \ s i = 0 \ foldseq f s n = s n" + "\x. f 0 x = x \ \i. i < n \ s i = 0 \ foldseq f s n = s n" proof - - assume a:"! i. i s i = 0" - assume x:"! x. f 0 x = x" - from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast - have b: "!! i l. i < Suc l = (i <= l)" by arith - have d: "!! k. k \ 0 \ ? l. k = Suc l" by arith + assume a: "\i. i s i = 0" + assume x: "\x. f 0 x = x" + from x have f00x: "\x. f 0 (f 0 x) = f 0 x" by blast + have b: "\i l. i < Suc l = (i <= l)" by arith + have d: "\k. k \ 0 \ \l. k = Suc l" by arith show "foldseq f s n = s n" apply (case_tac "n=0") apply (simp) @@ -610,11 +610,11 @@ qed lemma foldseq_almostzero: - assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \ j \ s i = 0" + assumes f0x: "\x. f 0 x = x" and fx0: "\x. f x 0 = x" and s0: "\i. i \ j \ s i = 0" shows "foldseq f s n = (if (j <= n) then (s j) else 0)" proof - - from s0 have a: "! i. i < j \ s i = 0" by simp - from s0 have b: "! i. j < i \ s i = 0" by simp + from s0 have a: "\i. i < j \ s i = 0" by simp + from s0 have b: "\i. j < i \ s i = 0" by simp show ?thesis apply auto apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym]) @@ -629,7 +629,7 @@ assumes "!! a b. g (f a b) = f (g a) (g b)" shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" proof - - have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n" + have "\s. g(foldseq f s n) = foldseq f (% x. g(s x)) n" apply (induct_tac n) apply (simp) apply (simp) @@ -668,10 +668,10 @@ qed definition r_distributive :: "('a \ 'b \ 'b) \ ('b \ 'b \ 'b) \ bool" where - "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" + "r_distributive fmul fadd == \a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" definition l_distributive :: "('a \ 'b \ 'a) \ ('a \ 'a \ 'a) \ bool" where - "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" + "l_distributive fmul fadd == \a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" definition distributive :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ bool" where "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd" @@ -685,8 +685,8 @@ "associative fadd" "commutative fadd" "fadd 0 0 = 0" - "! a. fmul a 0 = 0" - "! a. fmul 0 a = 0" + "\a. fmul a 0 = 0" + "\a. fmul 0 a = 0" shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" proof - from assms show ?thesis @@ -725,8 +725,8 @@ "associative fadd" "commutative fadd" "fadd 0 0 = 0" - "! a. fmul a 0 = 0" - "! a. fmul 0 a = 0" + "\a. fmul a 0 = 0" + "\a. fmul 0 a = 0" shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" proof - from assms show ?thesis @@ -796,20 +796,20 @@ apply (auto) by (subst foldseq_zero, (simp add: zero_matrix_def)+)+ -lemma mult_matrix_n_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix_n n fmul fadd A 0 = 0" +lemma mult_matrix_n_zero_right[simp]: "\fadd 0 0 = 0; \a. fmul a 0 = 0\ \ mult_matrix_n n fmul fadd A 0 = 0" apply (simp add: mult_matrix_n_def) apply (subst foldseq_zero) by (simp_all add: zero_matrix_def) -lemma mult_matrix_n_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix_n n fmul fadd 0 A = 0" +lemma mult_matrix_n_zero_left[simp]: "\fadd 0 0 = 0; \a. fmul 0 a = 0\ \ mult_matrix_n n fmul fadd 0 A = 0" apply (simp add: mult_matrix_n_def) apply (subst foldseq_zero) by (simp_all add: zero_matrix_def) -lemma mult_matrix_zero_left[simp]: "\fadd 0 0 = 0; !a. fmul 0 a = 0\ \ mult_matrix fmul fadd 0 A = 0" +lemma mult_matrix_zero_left[simp]: "\fadd 0 0 = 0; \a. fmul 0 a = 0\ \ mult_matrix fmul fadd 0 A = 0" by (simp add: mult_matrix_def) -lemma mult_matrix_zero_right[simp]: "\fadd 0 0 = 0; !a. fmul a 0 = 0\ \ mult_matrix fmul fadd A 0 = 0" +lemma mult_matrix_zero_right[simp]: "\fadd 0 0 = 0; \a. fmul a 0 = 0\ \ mult_matrix fmul fadd A 0 = 0" by (simp add: mult_matrix_def) lemma apply_matrix_zero[simp]: "f 0 = 0 \ apply_matrix f 0 = 0" @@ -987,10 +987,10 @@ lemma mult_matrix_singleton_right[simp]: assumes - "! x. fmul x 0 = 0" - "! x. fmul 0 x = 0" - "! x. fadd 0 x = x" - "! x. fadd x 0 = x" + "\x. fmul x 0 = 0" + "\x. fmul 0 x = 0" + "\x. fadd 0 x = x" + "\x. fadd x 0 = x" shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (% x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))" apply (simp add: mult_matrix_def) apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"]) @@ -1006,25 +1006,25 @@ lemma mult_matrix_ext: assumes eprem: - "? e. (! a b. a \ b \ fmul a e \ fmul b e)" + "\e. (\a b. a \ b \ fmul a e \ fmul b e)" and fprems: - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" - "! a. fadd a 0 = a" - "! a. fadd 0 a = a" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" + "\a. fadd a 0 = a" + "\a. fadd 0 a = a" and contraprems: "mult_matrix fmul fadd A = mult_matrix fmul fadd B" shows "A = B" proof(rule contrapos_np[of "False"], simp) assume a: "A \ B" - have b: "!! f g. (! x y. f x y = g x y) \ f = g" by ((rule ext)+, auto) - have "? j i. (Rep_matrix A j i) \ (Rep_matrix B j i)" + have b: "\f g. (\x y. f x y = g x y) \ f = g" by ((rule ext)+, auto) + have "\j i. (Rep_matrix A j i) \ (Rep_matrix B j i)" apply (rule contrapos_np[of "False"], simp+) apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp) by (simp add: Rep_matrix_inject a) then obtain J I where c:"(Rep_matrix A J I) \ (Rep_matrix B J I)" by blast - from eprem obtain e where eprops:"(! a b. a \ b \ fmul a e \ fmul b e)" by blast + from eprem obtain e where eprops:"(\a b. a \ b \ fmul a e \ fmul b e)" by blast let ?S = "singleton_matrix I 0 e" let ?comp = "mult_matrix fmul fadd" have d: "!!x f g. f = g \ f x = g x" by blast @@ -1046,12 +1046,12 @@ lemma foldmatrix_transpose: assumes - "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" + "\a b c d. g(f a b) (f c d) = f (g a c) (g b d)" shows "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" proof - - have forall:"!! P x. (! x. P x) \ P x" by auto - have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" + have forall:"\P x. (\x. P x) \ P x" by auto + have tworows:"\A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" apply (induct n) apply (simp add: foldmatrix_def foldmatrix_transposed_def assms)+ apply (auto) @@ -1069,7 +1069,7 @@ assumes "associative f" "associative g" - "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" + "\a b c d. g(f a b) (f c d) = f (g a c) (g b d)" shows "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n" apply (insert foldmatrix_transpose[of g f A m n]) @@ -1077,8 +1077,8 @@ lemma mult_n_nrows: assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" +"\a. fmul 0 a = 0" +"\a. fmul a 0 = 0" "fadd 0 0 = 0" shows "nrows (mult_matrix_n n fmul fadd A B) \ nrows A" apply (subst nrows_le) @@ -1093,8 +1093,8 @@ lemma mult_n_ncols: assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" +"\a. fmul 0 a = 0" +"\a. fmul a 0 = 0" "fadd 0 0 = 0" shows "ncols (mult_matrix_n n fmul fadd A B) \ ncols B" apply (subst ncols_le) @@ -1109,16 +1109,16 @@ lemma mult_nrows: assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" +"\a. fmul 0 a = 0" +"\a. fmul a 0 = 0" "fadd 0 0 = 0" shows "nrows (mult_matrix fmul fadd A B) \ nrows A" by (simp add: mult_matrix_def mult_n_nrows assms) lemma mult_ncols: assumes -"! a. fmul 0 a = 0" -"! a. fmul a 0 = 0" +"\a. fmul 0 a = 0" +"\a. fmul a 0 = 0" "fadd 0 0 = 0" shows "ncols (mult_matrix fmul fadd A B) \ ncols B" by (simp add: mult_matrix_def mult_n_ncols assms) @@ -1137,18 +1137,18 @@ lemma mult_matrix_assoc: assumes - "! a. fmul1 0 a = 0" - "! a. fmul1 a 0 = 0" - "! a. fmul2 0 a = 0" - "! a. fmul2 a 0 = 0" + "\a. fmul1 0 a = 0" + "\a. fmul1 a 0 = 0" + "\a. fmul2 0 a = 0" + "\a. fmul2 a 0 = 0" "fadd1 0 0 = 0" "fadd2 0 0 = 0" - "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" + "\a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" "associative fadd1" "associative fadd2" - "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" - "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" - "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" + "\a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" + "\a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" + "\a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" proof - have comb_left: "!! A B x y. A = B \ (Rep_matrix (Abs_matrix A)) x y = (Rep_matrix(Abs_matrix B)) x y" by blast @@ -1195,18 +1195,18 @@ lemma assumes - "! a. fmul1 0 a = 0" - "! a. fmul1 a 0 = 0" - "! a. fmul2 0 a = 0" - "! a. fmul2 a 0 = 0" + "\a. fmul1 0 a = 0" + "\a. fmul1 a 0 = 0" + "\a. fmul2 0 a = 0" + "\a. fmul2 a 0 = 0" "fadd1 0 0 = 0" "fadd2 0 0 = 0" - "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" + "\a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" "associative fadd1" "associative fadd2" - "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" - "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" - "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" + "\a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" + "\a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" + "\a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" shows "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)" apply (rule ext)+ @@ -1216,8 +1216,8 @@ lemma mult_matrix_assoc_simple: assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" "fadd 0 0 = 0" "associative fadd" "commutative fadd" @@ -1247,8 +1247,8 @@ lemma Rep_mult_matrix: assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" "fadd 0 0 = 0" shows "Rep_matrix(mult_matrix fmul fadd A B) j i = @@ -1262,10 +1262,10 @@ lemma transpose_mult_matrix: assumes - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" "fadd 0 0 = 0" - "! x y. fmul y x = fmul x y" + "\x y. fmul y x = fmul x y" shows "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)" apply (simp add: Rep_matrix_inject[THEN sym]) @@ -1314,7 +1314,7 @@ lemma le_apply_matrix: assumes "f 0 = 0" - "! x y. x <= y \ f x <= f y" + "\x y. x <= y \ f x <= f y" "(a::('a::{ord, zero}) matrix) <= b" shows "apply_matrix f a <= apply_matrix f b" @@ -1323,7 +1323,7 @@ lemma le_combine_matrix: assumes "f 0 0 = 0" - "! a b c d. a <= b & c <= d \ f a c <= f b d" + "\a b c d. a <= b & c <= d \ f a c <= f b d" "A <= B" "C <= D" shows @@ -1333,7 +1333,7 @@ lemma le_left_combine_matrix: assumes "f 0 0 = 0" - "! a b c. a <= b \ f c a <= f c b" + "\a b c. a <= b \ f c a <= f c b" "A <= B" shows "combine_matrix f C A <= combine_matrix f C B" @@ -1342,7 +1342,7 @@ lemma le_right_combine_matrix: assumes "f 0 0 = 0" - "! a b c. a <= b \ f a c <= f b c" + "\a b c. a <= b \ f a c <= f b c" "A <= B" shows "combine_matrix f A C <= combine_matrix f B C" @@ -1353,22 +1353,22 @@ lemma le_foldseq: assumes - "! a b c d . a <= b & c <= d \ f a c <= f b d" - "! i. i <= n \ s i <= t i" + "\a b c d . a <= b & c <= d \ f a c <= f b d" + "\i. i <= n \ s i <= t i" shows "foldseq f s n <= foldseq f t n" proof - - have "! s t. (! i. i<=n \ s i <= t i) \ foldseq f s n <= foldseq f t n" + have "\s t. (\i. i<=n \ s i <= t i) \ foldseq f s n <= foldseq f t n" by (induct n) (simp_all add: assms) then show "foldseq f s n <= foldseq f t n" using assms by simp qed lemma le_left_mult: assumes - "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" - "! c a b. 0 <= c & a <= b \ fmul c a <= fmul c b" - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" + "\a b c d. a <= b & c <= d \ fadd a c <= fadd b d" + "\c a b. 0 <= c & a <= b \ fmul c a <= fmul c b" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" "fadd 0 0 = 0" "0 <= C" "A <= B" @@ -1384,10 +1384,10 @@ lemma le_right_mult: assumes - "! a b c d. a <= b & c <= d \ fadd a c <= fadd b d" - "! c a b. 0 <= c & a <= b \ fmul a c <= fmul b c" - "! a. fmul 0 a = 0" - "! a. fmul a 0 = 0" + "\a b c d. a <= b & c <= d \ fadd a c <= fadd b d" + "\c a b. 0 <= c & a <= b \ fmul a c <= fmul b c" + "\a. fmul 0 a = 0" + "\a. fmul a 0 = 0" "fadd 0 0 = 0" "0 <= C" "A <= B" @@ -1401,7 +1401,7 @@ apply (auto) done -lemma spec2: "! j i. P j i \ P j i" by blast +lemma spec2: "\j i. P j i \ P j i" by blast lemma neg_imp: "(\ Q \ \ P) \ P \ Q" by blast lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))" @@ -1628,7 +1628,7 @@ apply (simp add: Rep_mult_matrix) done -lemma apply_matrix_add: "! x y. f (x+y) = (f x) + (f y) \ f 0 = (0::'a) +lemma apply_matrix_add: "\x y. f (x+y) = (f x) + (f y) \ f 0 = (0::'a) \ apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)" apply (subst Rep_matrix_inject[symmetric]) apply (rule ext)+ @@ -1751,7 +1751,7 @@ by auto lemma Rep_matrix_zero_imp_mult_zero: - "! j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0 \ A * B = (0::('a::lattice_ring) matrix)" + "\j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0 \ A * B = (0::('a::lattice_ring) matrix)" apply (subst Rep_matrix_inject[symmetric]) apply (rule ext)+ apply (auto simp add: Rep_matrix_mult foldseq_zero zero_imp_mult_zero) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Matrix_LP/SparseMatrix.thy --- a/src/HOL/Matrix_LP/SparseMatrix.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Matrix_LP/SparseMatrix.thy Thu Feb 15 13:04:36 2018 +0100 @@ -25,7 +25,7 @@ lemmas [code] = sparse_row_vector_empty [symmetric] -lemma foldl_distrstart: "! a x y. (f (g x y) a = g x (f y a)) \ (foldl f (g x y) l = g x (foldl f y l))" +lemma foldl_distrstart: "\a x y. (f (g x y) a = g x (f y a)) \ (foldl f (g x y) l = g x (foldl f y l))" by (induct l arbitrary: x y, auto) lemma sparse_row_vector_cons[simp]: @@ -180,7 +180,7 @@ lemma addmult_spvec_empty2[simp]: "addmult_spvec y a [] = a" by (induct a) auto -lemma sparse_row_vector_map: "(! x y. f (x+y) = (f x) + (f y)) \ (f::'a\('a::lattice_ring)) 0 = 0 \ +lemma sparse_row_vector_map: "(\x y. f (x+y) = (f x) + (f y)) \ (f::'a\('a::lattice_ring)) 0 = 0 \ sparse_row_vector (map (% x. (fst x, f (snd x))) a) = apply_matrix f (sparse_row_vector a)" apply (induct a) apply (simp_all add: apply_matrix_add) @@ -417,7 +417,7 @@ lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" proof - - have "(! x ab a. x = (a,b)#arr \ add_spvec x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" + have "(\x ab a. x = (a,b)#arr \ add_spvec x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" by (induct brr rule: add_spvec.induct) (auto split:if_splits) then show ?thesis by (case_tac brr, auto) @@ -425,7 +425,7 @@ lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr) brr = (ab, bb) # list \ (ab = a | (brr \ [] & ab = fst (hd brr)))" proof - - have "(! x ab a. x = (a,b)#arr \ add_spmat x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" + have "(\x ab a. x = (a,b)#arr \ add_spmat x brr = (ab, bb) # list \ (ab = a | (ab = fst (hd brr))))" by (rule add_spmat.induct) (auto split:if_splits) then show ?thesis by (case_tac brr, auto) @@ -559,7 +559,7 @@ definition disj_matrices :: "('a::zero) matrix \ 'a matrix \ bool" where "disj_matrices A B \ - (! j i. (Rep_matrix A j i \ 0) \ (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \ 0) \ (Rep_matrix A j i = 0))" + (\j i. (Rep_matrix A j i \ 0) \ (Rep_matrix B j i = 0)) & (\j i. (Rep_matrix B j i \ 0) \ (Rep_matrix A j i = 0))" declare [[simp_depth_limit = 6]] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Metis_Examples/Big_O.thy --- a/src/HOL/Metis_Examples/Big_O.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Metis_Examples/Big_O.thy Thu Feb 15 13:04:36 2018 +0100 @@ -128,17 +128,17 @@ lemma bigo_alt_def: "O(f) = {h. \c. (0 < c \ (\x. \h x\ <= c * \f x\))}" by (auto simp add: bigo_def bigo_pos_const) -lemma bigo_elt_subset [intro]: "f : O(g) \ O(f) \ O(g)" +lemma bigo_elt_subset [intro]: "f \ O(g) \ O(f) \ O(g)" apply (auto simp add: bigo_alt_def) apply (rule_tac x = "ca * c" in exI) apply (metis algebra_simps mult_le_cancel_left_pos order_trans mult_pos_pos) done -lemma bigo_refl [intro]: "f : O(f)" +lemma bigo_refl [intro]: "f \ O(f)" unfolding bigo_def mem_Collect_eq by (metis mult_1 order_refl) -lemma bigo_zero: "0 : O(g)" +lemma bigo_zero: "0 \ O(g)" apply (auto simp add: bigo_def func_zero) by (metis mult_zero_left order_refl) @@ -210,12 +210,12 @@ apply (metis max.cobounded2 linorder_linear max.absorb1 mult_right_mono xt1(6)) by (metis max.cobounded2 linorder_not_le mult_le_cancel_right order_trans) -lemma bigo_bounded_alt: "\x. 0 <= f x \ \x. f x <= c * g x \ f : O(g)" +lemma bigo_bounded_alt: "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" apply (auto simp add: bigo_def) (* Version 1: one-line proof *) by (metis abs_le_D1 linorder_class.not_less order_less_le Orderings.xt1(12) abs_mult) -lemma "\x. 0 <= f x \ \x. f x <= c * g x \ f : O(g)" +lemma "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" apply (auto simp add: bigo_def) (* Version 2: structured proof *) proof - @@ -223,11 +223,11 @@ thus "\c. \x. f x \ c * \g x\" by (metis abs_mult abs_ge_self order_trans) qed -lemma bigo_bounded: "\x. 0 <= f x \ \x. f x <= g x \ f : O(g)" +lemma bigo_bounded: "\x. 0 \ f x \ \x. f x \ g x \ f \ O(g)" apply (erule bigo_bounded_alt [of f 1 g]) by (metis mult_1) -lemma bigo_bounded2: "\x. lb x <= f x \ \x. f x <= lb x + g x \ f : lb +o O(g)" +lemma bigo_bounded2: "\x. lb x \ f x \ \x. f x \ lb x + g x \ f \ lb +o O(g)" apply (rule set_minus_imp_plus) apply (rule bigo_bounded) apply (metis add_le_cancel_left diff_add_cancel diff_self minus_apply @@ -258,40 +258,40 @@ apply (rule set_minus_imp_plus) apply (subst fun_diff_def) proof - - assume a: "f - g : O(h)" + assume a: "f - g \ O(h)" have "(\x. \f x\ - \g x\) =o O(\x. \\f x\ - \g x\\)" by (rule bigo_abs2) - also have "... <= O(\x. \f x - g x\)" + also have "\ <= O(\x. \f x - g x\)" apply (rule bigo_elt_subset) apply (rule bigo_bounded) apply (metis abs_ge_zero) by (metis abs_triangle_ineq3) - also have "... <= O(f - g)" + also have "\ <= O(f - g)" apply (rule bigo_elt_subset) apply (subst fun_diff_def) apply (rule bigo_abs) done - also have "... <= O(h)" + also have "\ <= O(h)" using a by (rule bigo_elt_subset) - finally show "(\x. \f x\ - \g x\) : O(h)" . + finally show "(\x. \f x\ - \g x\) \ O(h)" . qed lemma bigo_abs5: "f =o O(g) \ (\x. \f x\) =o O(g)" by (unfold bigo_def, auto) -lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \ O(f) <= O(g) + O(h)" +lemma bigo_elt_subset2 [intro]: "f \ g +o O(h) \ O(f) \ O(g) + O(h)" proof - - assume "f : g +o O(h)" - also have "... <= O(g) + O(h)" + assume "f \ g +o O(h)" + also have "\ \ O(g) + O(h)" by (auto del: subsetI) - also have "... = O(\x. \g x\) + O(\x. \h x\)" + also have "\ = O(\x. \g x\) + O(\x. \h x\)" by (metis bigo_abs3) also have "... = O((\x. \g x\) + (\x. \h x\))" by (rule bigo_plus_eq [symmetric], auto) - finally have "f : ...". - then have "O(f) <= ..." + finally have "f \ \". + then have "O(f) \ \" by (elim bigo_elt_subset) - also have "... = O(\x. \g x\) + O(\x. \h x\)" + also have "\ = O(\x. \g x\) + O(\x. \h x\)" by (rule bigo_plus_eq, auto) finally show ?thesis by (simp add: bigo_abs3 [symmetric]) @@ -313,32 +313,32 @@ lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" by (metis bigo_mult bigo_refl set_times_mono3 subset_trans) -lemma bigo_mult3: "f : O(h) \ g : O(j) \ f * g : O(h * j)" +lemma bigo_mult3: "f \ O(h) \ g \ O(j) \ f * g \ O(h * j)" by (metis bigo_mult set_rev_mp set_times_intro) -lemma bigo_mult4 [intro]:"f : k +o O(h) \ g * f : (g * k) +o O(g * h)" +lemma bigo_mult4 [intro]:"f \ k +o O(h) \ g * f \ (g * k) +o O(g * h)" by (metis bigo_mult2 set_plus_mono_b set_times_intro2 set_times_plus_distrib) lemma bigo_mult5: "\x. f x ~= 0 \ O(f * g) <= (f::'a => ('b::linordered_field)) *o O(g)" proof - - assume a: "\x. f x ~= 0" + assume a: "\x. f x \ 0" show "O(f * g) <= f *o O(g)" proof fix h - assume h: "h : O(f * g)" - then have "(\x. 1 / (f x)) * h : (\x. 1 / f x) *o O(f * g)" + assume h: "h \ O(f * g)" + then have "(\x. 1 / (f x)) * h \ (\x. 1 / f x) *o O(f * g)" by auto also have "... <= O((\x. 1 / f x) * (f * g))" by (rule bigo_mult2) also have "(\x. 1 / f x) * (f * g) = g" by (simp add: fun_eq_iff a) - finally have "(\x. (1::'b) / f x) * h : O(g)". - then have "f * ((\x. (1::'b) / f x) * h) : f *o O(g)" + finally have "(\x. (1::'b) / f x) * h \ O(g)". + then have "f * ((\x. (1::'b) / f x) * h) \ f *o O(g)" by auto also have "f * ((\x. (1::'b) / f x) * h) = h" by (simp add: func_times fun_eq_iff a) - finally show "h : f *o O(g)". + finally show "h \ f *o O(g)". qed qed @@ -360,61 +360,61 @@ "\x. f x \ 0 \ O(f * g) = O(f::'a \ ('b::linordered_field)) * O(g)" by (metis bigo_mult bigo_mult7 order_antisym_conv) -lemma bigo_minus [intro]: "f : O(g) \ - f : O(g)" +lemma bigo_minus [intro]: "f \ O(g) \ - f \ O(g)" by (auto simp add: bigo_def fun_Compl_def) -lemma bigo_minus2: "f : g +o O(h) \ -f : -g +o O(h)" +lemma bigo_minus2: "f \ g +o O(h) \ -f \ -g +o O(h)" by (metis (no_types, lifting) bigo_minus diff_minus_eq_add minus_add_distrib minus_minus set_minus_imp_plus set_plus_imp_minus) lemma bigo_minus3: "O(-f) = O(f)" by (metis bigo_elt_subset bigo_minus bigo_refl equalityI minus_minus) -lemma bigo_plus_absorb_lemma1: "f : O(g) \ f +o O(g) \ O(g)" +lemma bigo_plus_absorb_lemma1: "f \ O(g) \ f +o O(g) \ O(g)" by (metis bigo_plus_idemp set_plus_mono3) -lemma bigo_plus_absorb_lemma2: "f : O(g) \ O(g) \ f +o O(g)" +lemma bigo_plus_absorb_lemma2: "f \ O(g) \ O(g) \ f +o O(g)" by (metis (no_types) bigo_minus bigo_plus_absorb_lemma1 right_minus set_plus_mono set_plus_rearrange2 set_zero_plus subsetD subset_refl subset_trans) -lemma bigo_plus_absorb [simp]: "f : O(g) \ f +o O(g) = O(g)" +lemma bigo_plus_absorb [simp]: "f \ O(g) \ f +o O(g) = O(g)" by (metis bigo_plus_absorb_lemma1 bigo_plus_absorb_lemma2 order_eq_iff) -lemma bigo_plus_absorb2 [intro]: "f : O(g) \ A <= O(g) \ f +o A \ O(g)" +lemma bigo_plus_absorb2 [intro]: "f \ O(g) \ A \ O(g) \ f +o A \ O(g)" by (metis bigo_plus_absorb set_plus_mono) -lemma bigo_add_commute_imp: "f : g +o O(h) \ g : f +o O(h)" +lemma bigo_add_commute_imp: "f \ g +o O(h) \ g \ f +o O(h)" by (metis bigo_minus minus_diff_eq set_plus_imp_minus set_minus_plus) -lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))" +lemma bigo_add_commute: "(f \ g +o O(h)) = (g \ f +o O(h))" by (metis bigo_add_commute_imp) -lemma bigo_const1: "(\x. c) : O(\x. 1)" +lemma bigo_const1: "(\x. c) \ O(\x. 1)" by (auto simp add: bigo_def ac_simps) lemma bigo_const2 [intro]: "O(\x. c) \ O(\x. 1)" by (metis bigo_const1 bigo_elt_subset) -lemma bigo_const3: "(c::'a::linordered_field) ~= 0 \ (\x. 1) : O(\x. c)" +lemma bigo_const3: "(c::'a::linordered_field) \ 0 \ (\x. 1) \ O(\x. c)" apply (simp add: bigo_def) by (metis abs_eq_0 left_inverse order_refl) -lemma bigo_const4: "(c::'a::linordered_field) ~= 0 \ O(\x. 1) <= O(\x. c)" +lemma bigo_const4: "(c::'a::linordered_field) \ 0 \ O(\x. 1) \ O(\x. c)" by (metis bigo_elt_subset bigo_const3) lemma bigo_const [simp]: "(c::'a::linordered_field) ~= 0 \ O(\x. c) = O(\x. 1)" by (metis bigo_const2 bigo_const4 equalityI) -lemma bigo_const_mult1: "(\x. c * f x) : O(f)" +lemma bigo_const_mult1: "(\x. c * f x) \ O(f)" apply (simp add: bigo_def abs_mult) by (metis le_less) lemma bigo_const_mult2: "O(\x. c * f x) \ O(f)" by (rule bigo_elt_subset, rule bigo_const_mult1) -lemma bigo_const_mult3: "(c::'a::linordered_field) ~= 0 \ f : O(\x. c * f x)" +lemma bigo_const_mult3: "(c::'a::linordered_field) \ 0 \ f \ O(\x. c * f x)" apply (simp add: bigo_def) by (metis (no_types) abs_mult mult.assoc mult_1 order_refl left_inverse) @@ -422,11 +422,11 @@ "(c::'a::linordered_field) \ 0 \ O(f) \ O(\x. c * f x)" by (metis bigo_elt_subset bigo_const_mult3) -lemma bigo_const_mult [simp]: "(c::'a::linordered_field) ~= 0 \ +lemma bigo_const_mult [simp]: "(c::'a::linordered_field) \ 0 \ O(\x. c * f x) = O(f)" by (metis equalityI bigo_const_mult2 bigo_const_mult4) -lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) ~= 0 \ +lemma bigo_const_mult5 [simp]: "(c::'a::linordered_field) \ 0 \ (\x. c) *o O(f) = O(f)" apply (auto del: subsetI) apply (rule order_trans) @@ -491,7 +491,7 @@ subsection \Sum\ -lemma bigo_sum_main: "\x. \y \ A x. 0 <= h x y \ +lemma bigo_sum_main: "\x. \y \ A x. 0 \ h x y \ \c. \x. \y \ A x. \f x y\ <= c * (h x y) \ (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" apply (auto simp add: bigo_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Metis_Examples/Clausification.thy --- a/src/HOL/Metis_Examples/Clausification.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Metis_Examples/Clausification.thy Thu Feb 15 13:04:36 2018 +0100 @@ -137,7 +137,7 @@ (\ys x zs. xs = ys @ x # zs \ P x \ (\z \ set zs. \ P z))" by (metis split_list_last_prop[where P = P] in_set_conv_decomp) -lemma ex_tl: "EX ys. tl ys = xs" +lemma ex_tl: "\ys. tl ys = xs" using list.sel(3) by fast lemma "(\ys::nat list. tl ys = xs) \ (\bs::int list. tl bs = as)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Metis_Examples/Message.thy --- a/src/HOL/Metis_Examples/Message.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Metis_Examples/Message.thy Thu Feb 15 13:04:36 2018 +0100 @@ -84,7 +84,7 @@ by (metis parts.Body) text\Equations hold because constructors are injective.\ -lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x:A)" +lemma Friend_image_eq [simp]: "(Friend x \ Friend`A) = (x\A)" by (metis agent.inject image_iff) lemma Key_image_eq [simp]: "(Key x \ Key`A) = (x \ A)" @@ -324,7 +324,7 @@ | 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" + "[|Crypt K X \ analz H; Key(invKey K) \ analz H|] ==> X \ analz H" text\Monotonicity; Lemma 1 of Lowe's paper\ @@ -723,7 +723,7 @@ qed lemma Fake_parts_insert_in_Un: - "[|Z \ parts (insert X H); X: synth (analz H)|] + "[|Z \ parts (insert X H); X \ synth (analz H)|] ==> Z \ synth (analz H) \ parts H" by (blast dest: Fake_parts_insert [THEN subsetD, dest]) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Metis_Examples/Sets.thy --- a/src/HOL/Metis_Examples/Sets.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Metis_Examples/Sets.thy Thu Feb 15 13:04:36 2018 +0100 @@ -13,12 +13,12 @@ declare [[metis_new_skolem]] -lemma "EX x X. ALL y. EX z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & +lemma "\x X. \y. \z Z. (~P(y,y) | P(x,x) | ~S(z,x)) & (S(x,y) | ~S(y,z) | Q(Z,Z)) & (Q(X,y) | ~Q(y,Z) | S(X,X))" by metis -lemma "P(n::nat) ==> ~P(0) ==> n ~= 0" +lemma "P(n::nat) ==> \P(0) ==> n \ 0" by metis sledgehammer_params [isar_proofs, compress = 1] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Metis_Examples/Tarski.thy Thu Feb 15 13:04:36 2018 +0100 @@ -23,35 +23,35 @@ order :: "('a * 'a) set" definition monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" where - "monotone f A r == \x\A. \y\A. (x, y): r --> ((f x), (f y)) : r" + "monotone f A r == \x\A. \y\A. (x, y) \ r --> ((f x), (f y)) \ r" definition least :: "['a => bool, 'a potype] => 'a" where - "least P po == @ x. x: pset po & P x & - (\y \ pset po. P y --> (x,y): order po)" + "least P po \ SOME x. x \ pset po \ P x \ + (\y \ pset po. P y \ (x,y) \ order po)" definition greatest :: "['a => bool, 'a potype] => 'a" where - "greatest P po == @ x. x: pset po & P x & - (\y \ pset po. P y --> (y,x): order po)" + "greatest P po \ SOME x. x \ pset po \ P x \ + (\y \ pset po. P y \ (y,x) \ order po)" definition lub :: "['a set, 'a potype] => 'a" where - "lub S po == least (\x. \y\S. (y,x): order po) po" + "lub S po == least (\x. \y\S. (y,x) \ order po) po" definition glb :: "['a set, 'a potype] => 'a" where - "glb S po == greatest (\x. \y\S. (x,y): order po) po" + "glb S po \ greatest (\x. \y\S. (x,y) \ order po) po" definition isLub :: "['a set, 'a potype, 'a] => bool" where - "isLub S po == \L. (L: pset po & (\y\S. (y,L): order po) & - (\z\pset po. (\y\S. (y,z): order po) --> (L,z): order po))" + "isLub S po \ \L. (L \ pset po \ (\y\S. (y,L) \ order po) \ + (\z\pset po. (\y\S. (y,z) \ order po) \ (L,z) \ order po))" definition isGlb :: "['a set, 'a potype, 'a] => bool" where - "isGlb S po == \G. (G: pset po & (\y\S. (G,y): order po) & - (\z \ pset po. (\y\S. (z,y): order po) --> (z,G): order po))" + "isGlb S po \ \G. (G \ pset po \ (\y\S. (G,y) \ order po) \ + (\z \ pset po. (\y\S. (z,y) \ order po) \ (z,G) \ order po))" definition "fix" :: "[('a => 'a), 'a set] => 'a set" where - "fix f A == {x. x: A & f x = x}" + "fix f A \ {x. x \ A \ f x = x}" definition interval :: "[('a*'a) set,'a, 'a ] => 'a set" where - "interval r a b == {x. (a,x): r & (x,b): r}" + "interval r a b == {x. (a,x) \ r & (x,b) \ r}" definition Bot :: "'a potype => 'a" where "Bot po == least (\x. True) po" @@ -64,22 +64,22 @@ trans (order P)}" definition CompleteLattice :: "('a potype) set" where - "CompleteLattice == {cl. cl: PartialOrder & - (\S. S \ pset cl --> (\L. isLub S cl L)) & - (\S. S \ pset cl --> (\G. isGlb S cl G))}" + "CompleteLattice == {cl. cl \ PartialOrder \ + (\S. S \ pset cl \ (\L. isLub S cl L)) \ + (\S. S \ pset cl \ (\G. isGlb S cl G))}" definition induced :: "['a set, ('a * 'a) set] => ('a *'a)set" where - "induced A r == {(a,b). a : A & b: A & (a,b): r}" + "induced A r \ {(a,b). a \ A \ b \ A \ (a,b) \ r}" definition sublattice :: "('a potype * 'a set)set" where - "sublattice == - SIGMA cl: CompleteLattice. - {S. S \ pset cl & - (| pset = S, order = induced S (order cl) |): CompleteLattice }" + "sublattice \ + SIGMA cl : CompleteLattice. + {S. S \ pset cl \ + \pset = S, order = induced S (order cl)\ \ CompleteLattice}" abbreviation sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50) - where "S <<= cl \ S : sublattice `` {cl}" + where "S <<= cl \ S \ sublattice `` {cl}" definition dual :: "'a potype => 'a potype" where "dual po == (| pset = pset po, order = converse (order po) |)" @@ -88,21 +88,21 @@ fixes cl :: "'a potype" and A :: "'a set" and r :: "('a * 'a) set" - assumes cl_po: "cl : PartialOrder" + assumes cl_po: "cl \ PartialOrder" defines A_def: "A == pset cl" and r_def: "r == order cl" locale CL = PO + - assumes cl_co: "cl : CompleteLattice" + assumes cl_co: "cl \ CompleteLattice" definition CLF_set :: "('a potype * ('a => 'a)) set" where "CLF_set = (SIGMA cl: CompleteLattice. - {f. f: pset cl \ pset cl & monotone f (pset cl) (order cl)})" + {f. f \ pset cl \ pset cl \ monotone f (pset cl) (order cl)})" locale CLF = CL + fixes f :: "'a => 'a" and P :: "'a set" - assumes f_cl: "(cl,f) : CLF_set" (*was the equivalent "f : CLF``{cl}"*) + assumes f_cl: "(cl,f) \ CLF_set" (*was the equivalent "f : CLF``{cl}"*) defines P_def: "P == fix f A" locale Tarski = CLF + @@ -113,9 +113,9 @@ Y_ss: "Y \ P" defines intY1_def: "intY1 == interval r (lub Y cl) (Top cl)" - and v_def: "v == glb {x. ((\x \ intY1. f x) x, x): induced intY1 r & - x: intY1} - (| pset=intY1, order=induced intY1 r|)" + and v_def: "v == glb {x. ((\x \ intY1. f x) x, x) \ induced intY1 r \ + x \ intY1} + \pset=intY1, order=induced intY1 r\" subsection \Partial Order\ @@ -362,8 +362,8 @@ by (simp add: isLub_def A_def r_def) lemma (in CL) isLubI: - "[| L \ A; \y \ S. (y, L) \ r; - (\z \ A. (\y \ S. (y, z):r) --> (L, z) \ r)|] ==> isLub S cl L" + "\L \ A; \y \ S. (y, L) \ r; + (\z \ A. (\y \ S. (y, z) \ r) \ (L, z) \ r)\ \ isLub S cl L" by (simp add: isLub_def A_def r_def) subsection \glb\ @@ -402,7 +402,7 @@ declare (in CLF) f_cl [simp] lemma (in CLF) [simp]: - "f: pset cl \ pset cl & monotone f (pset cl) (order cl)" + "f \ pset cl \ pset cl \ monotone f (pset cl) (order cl)" proof - have "\u v. (v, u) \ CLF_set \ u \ {R \ pset v \ pset v. monotone R (pset v) (order v)}" unfolding CLF_set_def using SigmaE2 by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Thu Feb 15 13:04:36 2018 +0100 @@ -433,7 +433,7 @@ qed definition some_elem :: "'a set \ 'a" where [code del]: - "some_elem = (%S. SOME x. x : S)" + "some_elem = (\S. SOME x. x \ S)" code_printing constant some_elem \ (SML) "(case/ _ of/ Set/ xs/ =>/ hd/ xs)" @@ -486,4 +486,3 @@ \ end - diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/BV/JType.thy --- a/src/HOL/MicroJava/BV/JType.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/BV/JType.thy Thu Feb 15 13:04:36 2018 +0100 @@ -33,7 +33,7 @@ definition is_ty :: "'c prog \ ty \ bool" where "is_ty G T == case T of PrimT P \ True | RefT R \ - (case R of NullT \ True | ClassT C \ (C, Object) \ (subcls1 G)^*)" + (case R of NullT \ True | ClassT C \ (C, Object) \ (subcls1 G)\<^sup>*)" abbreviation "types G == Collect (is_type G)" @@ -102,16 +102,16 @@ done lemma wf_converse_subcls1_impl_acc_subtype: - "wf ((subcls1 G)^-1) \ acc (subtype G)" + "wf ((subcls1 G)\) \ acc (subtype G)" apply (unfold Semilat.acc_def lesssub_def) -apply (drule_tac p = "((subcls1 G)^-1) - Id" in wf_subset) +apply (drule_tac p = "((subcls1 G)\) - Id" in wf_subset) apply auto apply (drule wf_trancl) apply (simp add: wf_eq_minimal) apply clarify apply (unfold lesub_def subtype_def) apply (rename_tac M T) -apply (case_tac "EX C. Class C : M") +apply (case_tac "\C. Class C \ M") prefer 2 apply (case_tac T) apply (fastforce simp add: PrimT_PrimT2) @@ -131,7 +131,7 @@ apply (case_tac ref_ty) apply simp apply simp -apply (erule_tac x = "{C. Class C : M}" in allE) +apply (erule_tac x = "{C. Class C \ M}" in allE) apply auto apply (rename_tac D) apply (rule_tac x = "Class D" in bexI) @@ -148,7 +148,7 @@ apply (erule rtrancl.cases) apply blast apply (drule rtrancl_converseI) -apply (subgoal_tac "(subcls1 G - Id)^-1 = (subcls1 G)^-1 - Id") +apply (subgoal_tac "(subcls1 G - Id)\ = (subcls1 G)\ - Id") prefer 2 apply (simp add: converse_Int) apply safe[1] apply simp @@ -185,7 +185,7 @@ by (blast intro: subcls_C_Object) with ws_prog single_valued obtain u where - "is_lub ((subcls1 G)^* ) c1 c2 u" + "is_lub ((subcls1 G)\<^sup>*) c1 c2 u" by (blast dest: single_valued_has_lubs) moreover note acyclic @@ -226,7 +226,7 @@ by (blast intro: subcls_C_Object) with ws_prog single_valued obtain u where - lub: "is_lub ((subcls1 G)^*) c1 c2 u" + lub: "is_lub ((subcls1 G)\<^sup>*) c1 c2 u" by (blast dest: single_valued_has_lubs) with acyclic have "exec_lub (subcls1 G) (super G) c1 c2 = u" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/BV/Typing_Framework_JVM.thy --- a/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy Thu Feb 15 13:04:36 2018 +0100 @@ -53,7 +53,7 @@ by (unfold bounded_def) (blast dest: check_boundedD) lemma special_ex_swap_lemma [iff]: - "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)" + "(\X. (\n. X = A n \ P n) & Q X) = (\n. Q(A n) \ P n)" by blast lemmas [iff del] = not_None_eq diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Thu Feb 15 13:04:36 2018 +0100 @@ -740,7 +740,7 @@ (* case NewC *) apply clarify - apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)^-1) *) + apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)\) *) apply (simp add: c_hupd_hp_invariant) apply (rule progression_one_step) apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/Comp/LemmasComp.thy --- a/src/HOL/MicroJava/Comp/LemmasComp.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Thu Feb 15 13:04:36 2018 +0100 @@ -185,18 +185,18 @@ lemma comp_class_rec: - "wf ((subcls1 G)^-1) \ + "wf ((subcls1 G)\) \ class_rec (comp G) C t f = class_rec G C t (\ C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')" apply (rule_tac a = C in wf_induct) apply assumption - apply (subgoal_tac "wf ((subcls1 (comp G))^-1)") + apply (subgoal_tac "wf ((subcls1 (comp G))\)") apply (subgoal_tac "(class G x = None) \ (\ D fs ms. (class G x = Some (D, fs, ms)))") apply (erule disjE) (* case class G x = None *) apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 - wfrec [where R="(subcls1 G)^-1", simplified]) + wfrec [where R="(subcls1 G)\", simplified]) apply (simp add: comp_class_None) (* case \ D fs ms. (class G x = Some (D, fs, ms)) *) @@ -220,11 +220,11 @@ apply (simp add: comp_subcls1) done -lemma comp_fields: "wf ((subcls1 G)^-1) \ +lemma comp_fields: "wf ((subcls1 G)\) \ fields (comp G,C) = fields (G,C)" by (simp add: fields_def comp_class_rec) -lemma comp_field: "wf ((subcls1 G)^-1) \ +lemma comp_field: "wf ((subcls1 G)\) \ field (comp G,C) = field (G,C)" by (simp add: TypeRel.field_def comp_fields) @@ -233,7 +233,7 @@ \fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2); \C fs ms r1 r2. (R r1 r2) \ (R (f1 C fs ms r1) (f2 C fs ms r2)) \ \ ((class G C) \ None) \ R (class_rec G C t1 f1) (class_rec G C t2 f2)" - apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *) + apply (frule wf_subcls1) (* establish wf ((subcls1 G)\) *) apply (rule_tac a = C in wf_induct) apply assumption apply (intro strip) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Err.thy --- a/src/HOL/MicroJava/DFA/Err.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Err.thy Thu Feb 15 13:04:36 2018 +0100 @@ -34,7 +34,7 @@ "sup f == lift2(%x y. OK(x +_f y))" definition err :: "'a set \ 'a err set" where -"err A == insert Err {x . ? y:A. x = OK y}" +"err A == insert Err {x . \y\A. x = OK y}" definition esl :: "'a sl \ 'a esl" where "esl == %(A,r,f). (A,r, %x y. OK(f x y))" @@ -68,7 +68,7 @@ by (simp add: lesub_def) lemma le_err_refl: - "!x. x <=_r x \ e <=_(Err.le r) e" + "\x. x <=_r x \ e <=_(Err.le r) e" apply (unfold lesub_def Err.le_def) apply (simp split: err.split) done @@ -110,18 +110,18 @@ by (simp add: unfold_lesub_err le_def split: err.split) lemma le_OK_conv [iff]: - "e <=_(le r) OK x = (? y. e = OK y & y <=_r x)" + "e <=_(le r) OK x = (\y. e = OK y & y <=_r x)" by (simp add: unfold_lesub_err le_def split: err.split) lemma OK_le_conv: - "OK x <=_(le r) e = (e = Err | (? y. e = OK y & x <=_r y))" + "OK x <=_(le r) e = (e = Err | (\y. e = OK y & x <=_r y))" by (simp add: unfold_lesub_err le_def split: err.split) lemma top_Err [iff]: "top (le r) Err" by (simp add: top_def) lemma OK_less_conv [rule_format, iff]: - "OK x <_(le r) e = (e=Err | (? y. e = OK y & x <_r y))" + "OK x <_(le r) e = (e=Err | (\y. e = OK y & x <_r y))" by (simp add: lesssub_def lesub_def le_def split: err.split) lemma not_Err_less [rule_format, iff]: @@ -152,24 +152,24 @@ apply (unfold acc_def lesub_def le_def lesssub_def) apply (simp add: wf_eq_minimal split: err.split) apply clarify -apply (case_tac "Err : Q") +apply (case_tac "Err \ Q") apply blast -apply (erule_tac x = "{a . OK a : Q}" in allE) +apply (erule_tac x = "{a . OK a \ Q}" in allE) apply (case_tac "x") apply fast apply blast done -lemma Err_in_err [iff]: "Err : err A" +lemma Err_in_err [iff]: "Err \ err A" by (simp add: err_def) -lemma Ok_in_err [iff]: "(OK x : err A) = (x:A)" +lemma Ok_in_err [iff]: "(OK x \ err A) = (x\A)" by (auto simp add: err_def) subsection \lift\ lemma lift_in_errI: - "\ e : err S; !x:S. e = OK x \ f x : err S \ \ lift f e : err S" + "\ e \ err S; \x\S. e = OK x \ f x \ err S \ \ lift f e \ err S" apply (unfold lift_def) apply (simp split: err.split) apply blast @@ -203,7 +203,7 @@ by (simp add: plussub_def Err.sup_def Err.lift2_def) lemma Err_sup_eq_OK_conv [iff]: - "(Err.sup f ex ey = OK z) = (? x y. ex = OK x & ey = OK y & f x y = z)" + "(Err.sup f ex ey = OK z) = (\x y. ex = OK x & ey = OK y & f x y = z)" apply (unfold Err.sup_def lift2_def plussub_def) apply (rule iffI) apply (simp split: err.split_asm) @@ -220,17 +220,17 @@ subsection \semilat (err A) (le r) f\ lemma semilat_le_err_Err_plus [simp]: - "\ x: err A; semilat(err A, le r, f) \ \ Err +_f x = Err" + "\ x \ err A; semilat(err A, le r, f) \ \ Err +_f x = Err" by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1] Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1]) lemma semilat_le_err_plus_Err [simp]: - "\ x: err A; semilat(err A, le r, f) \ \ x +_f Err = Err" + "\ x \ err A; semilat(err A, le r, f) \ \ x +_f Err = Err" by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1] Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1]) lemma semilat_le_err_OK1: - "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ + "\ x \ A; y \ A; semilat(err A, le r, f); OK x +_f OK y = OK z \ \ x <=_r z" apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst) @@ -238,7 +238,7 @@ done lemma semilat_le_err_OK2: - "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ + "\ x \ A; y \ A; semilat(err A, le r, f); OK x +_f OK y = OK z \ \ y <=_r z" apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst) @@ -252,11 +252,11 @@ done lemma OK_plus_OK_eq_Err_conv [simp]: - assumes "x:A" and "y:A" and "semilat(err A, le r, fe)" - shows "((OK x) +_fe (OK y) = Err) = (~(? z:A. x <=_r z & y <=_r z))" + assumes "x \ A" and "y \ A" and "semilat(err A, le r, fe)" + shows "((OK x) +_fe (OK y) = Err) = (\(\z\A. x <=_r z & y <=_r z))" proof - have plus_le_conv3: "\A x y z f r. - \ semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \ + \ semilat (A,r,f); x +_f y <=_r z; x \ A; y \ A; z \ A \ \ x <=_r z \ y <=_r z" by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1]) from assms show ?thesis @@ -274,7 +274,7 @@ apply (case_tac "(OK x) +_fe (OK y)") apply assumption apply (rename_tac z) - apply (subgoal_tac "OK z: err A") + apply (subgoal_tac "OK z \ err A") apply (drule eq_order_le) apply (erule Semilat.orderI [OF Semilat.intro]) apply (blast dest: plus_le_conv3) @@ -287,12 +287,12 @@ (* FIXME? *) lemma all_bex_swap_lemma [iff]: - "(!x. (? y:A. x = f y) \ P x) = (!y:A. P(f y))" + "(\x. (\y\A. x = f y) \ P x) = (\y\A. P(f y))" by blast lemma closed_err_Union_lift2I: - "\ !A:AS. closed (err A) (lift2 f); AS ~= {}; - !A:AS.!B:AS. A~=B \ (!a:A.!b:B. a +_f b = Err) \ + "\ \A\AS. closed (err A) (lift2 f); AS \ {}; + \A\AS. \B\AS. A\B \ (\a\A. \b\B. a +_f b = Err) \ \ closed (err (\AS)) (lift2 f)" apply (unfold closed_def err_def) apply simp @@ -307,8 +307,8 @@ which may not hold \ lemma err_semilat_UnionI: - "\ !A:AS. err_semilat(A, r, f); AS ~= {}; - !A:AS.!B:AS. A~=B \ (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \ + "\ \A\AS. err_semilat(A, r, f); AS \ {}; + \A\AS. \B\AS. A\B \ (\a\A. \b\B. \ a <=_r b & a +_f b = Err) \ \ err_semilat (\AS, r, f)" apply (unfold semilat_def sl_def) apply (simp add: closed_err_Union_lift2I) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Kildall.thy --- a/src/HOL/MicroJava/DFA/Kildall.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Kildall.thy Thu Feb 15 13:04:36 2018 +0100 @@ -203,7 +203,7 @@ q \ w \ q = p \ \ stable r step (merges f (step p (ss!p)) ss) q" apply (unfold stable_def) - apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ s' : A") + apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ s' \ A") prefer 2 apply clarify apply (erule pres_typeD) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Listn.thy --- a/src/HOL/MicroJava/DFA/Listn.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Listn.thy Thu Feb 15 13:04:36 2018 +0100 @@ -94,7 +94,7 @@ done lemma le_list_refl: - "!x. x <=_r x \ xs <=[r] xs" + "\x. x <=_r x \ xs <=[r] xs" apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done @@ -154,23 +154,23 @@ done lemma listI: - "\ length xs = n; set xs <= A \ \ xs : list n A" + "\ length xs = n; set xs <= A \ \ xs \ list n A" apply (unfold list_def) apply blast done lemma listE_length [simp]: - "xs : list n A \ length xs = n" + "xs \ list n A \ length xs = n" apply (unfold list_def) apply blast done lemma less_lengthI: - "\ xs : list n A; p < n \ \ p < length xs" + "\ xs \ list n A; p < n \ \ p < length xs" by simp lemma listE_set [simp]: - "xs : list n A \ set xs <= A" + "xs \ list n A \ set xs <= A" apply (unfold list_def) apply blast done @@ -182,19 +182,19 @@ done lemma in_list_Suc_iff: - "(xs : list (Suc n) A) = (\y\ A. \ys\ list n A. xs = y#ys)" + "(xs \ list (Suc n) A) = (\y\ A. \ys\ list n A. xs = y#ys)" apply (unfold list_def) apply (case_tac "xs") apply auto done lemma Cons_in_list_Suc [iff]: - "(x#xs : list (Suc n) A) = (x\ A & xs : list n A)" + "(x#xs \ list (Suc n) A) = (x\ A & xs \ list n A)" apply (simp add: in_list_Suc_iff) done lemma list_not_empty: - "\a. a\ A \ \xs. xs : list n A" + "\a. a\ A \ \xs. xs \ list n A" apply (induct "n") apply simp apply (simp add: in_list_Suc_iff) @@ -203,14 +203,14 @@ lemma nth_in [rule_format, simp]: - "!i n. length xs = n \ set xs <= A \ i < n \ (xs!i) : A" + "\i n. length xs = n \ set xs <= A \ i < n \ (xs!i) \ A" apply (induct "xs") apply simp apply (simp add: nth_Cons split: nat.split) done lemma listE_nth_in: - "\ xs : list n A; i < n \ \ (xs!i) : A" + "\ xs \ list n A; i < n \ \ (xs!i) \ A" by auto @@ -245,7 +245,7 @@ lemma listt_update_in_list [simp, intro!]: - "\ xs : list n A; x\ A \ \ xs[i := x] : list n A" + "\ xs \ list n A; x\ A \ \ xs[i := x] \ list n A" apply (unfold list_def) apply simp done @@ -261,7 +261,7 @@ by (simp add: plussub_def map2_def split: list.split) lemma length_plus_list [rule_format, simp]: - "!ys. length(xs +[f] ys) = min(length xs) (length ys)" + "\ys. length(xs +[f] ys) = min(length xs) (length ys)" apply (induct xs) apply simp apply clarify @@ -269,7 +269,7 @@ done lemma nth_plus_list [rule_format, simp]: - "!xs ys i. length xs = n \ length ys = n \ i + "\xs ys i. length xs = n \ length ys = n \ i (xs +[f] ys)!i = (xs!i) +_f (ys!i)" apply (induct n) apply simp @@ -295,7 +295,7 @@ done lemma (in Semilat) plus_list_lub [rule_format]: -shows "!xs ys zs. set xs <= A \ set ys <= A \ set zs <= A +shows "\xs ys zs. set xs <= A \ set ys <= A \ set zs <= A \ size xs = n & size ys = n \ xs <=[r] zs & ys <=[r] zs \ xs +[f] ys <=[r] zs" apply (unfold unfold_lesub_list) @@ -304,7 +304,7 @@ lemma (in Semilat) list_update_incr [rule_format]: "x\ A \ set xs <= A \ - (!i. i xs <=[r] xs[i := x +_f xs!i])" + (\i. i xs <=[r] xs[i := x +_f xs!i])" apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) apply (induct xs) @@ -342,7 +342,7 @@ apply (erule thin_rl) apply (erule thin_rl) apply blast -apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) +apply (erule_tac x = "{a. \xs. size xs = k \ a#xs \ M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) @@ -394,7 +394,7 @@ by(simp add: Listn_sl_aux split_tupled_all) lemma coalesce_in_err_list [rule_format]: - "!xes. xes : list n (err A) \ coalesce xes : err(list n A)" + "\xes. xes \ list n (err A) \ coalesce xes \ err(list n A)" apply (induct n) apply simp apply clarify @@ -409,8 +409,8 @@ lemma coalesce_eq_OK1_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs. coalesce (xs +[f] ys) = OK zs \ xs <=[r] zs))" + \xs. xs \ list n A \ (\ys. ys \ list n A \ + (\zs. coalesce (xs +[f] ys) = OK zs \ xs <=[r] zs))" apply (induct n) apply simp apply clarify @@ -422,8 +422,8 @@ lemma coalesce_eq_OK2_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs. coalesce (xs +[f] ys) = OK zs \ ys <=[r] zs))" + \xs. xs \ list n A \ (\ys. ys \ list n A \ + (\zs. coalesce (xs +[f] ys) = OK zs \ ys <=[r] zs))" apply (induct n) apply simp apply clarify @@ -447,9 +447,9 @@ lemma coalesce_eq_OK_ub_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs us. coalesce (xs +[f] ys) = OK zs & xs <=[r] us & ys <=[r] us - & us : list n A \ zs <=[r] us))" + \xs. xs \ list n A \ (\ys. ys \ list n A \ + (\zs us. coalesce (xs +[f] ys) = OK zs \ xs <=[r] us \ ys <=[r] us + \ us \ list n A \ zs <=[r] us))" apply (induct n) apply simp apply clarify @@ -470,9 +470,9 @@ lemma coalesce_eq_Err_D [rule_format]: "\ semilat(err A, Err.le r, lift2 f) \ - \ !xs. xs\ list n A \ (!ys. ys\ list n A \ + \ \xs. xs \ list n A \ (\ys. ys \ list n A \ coalesce (xs +[f] ys) = Err \ - ~(\zs\ list n A. xs <=[r] zs & ys <=[r] zs))" + \(\zs\ list n A. xs <=[r] zs \ ys <=[r] zs))" apply (induct n) apply simp apply clarify @@ -483,15 +483,15 @@ done lemma closed_err_lift2_conv: - "closed (err A) (lift2 f) = (\x\ A. \y\ A. x +_f y : err A)" + "closed (err A) (lift2 f) = (\x\ A. \y\ A. x +_f y \ err A)" apply (unfold closed_def) apply (simp add: err_def) done lemma closed_map2_list [rule_format]: "closed (err A) (lift2 f) \ - \xs. xs : list n A \ (\ys. ys : list n A \ - map2 f xs ys : list n (err A))" + \xs. xs \ list n A \ (\ys. ys \ list n A \ + map2 f xs ys \ list n (err A))" apply (unfold map2_def) apply (induct n) apply simp diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Opt.thy --- a/src/HOL/MicroJava/DFA/Opt.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Opt.thy Thu Feb 15 13:04:36 2018 +0100 @@ -15,7 +15,7 @@ | Some x \ x <=_r y)" definition opt :: "'a set \ 'a option set" where -"opt A == insert None {x . ? y:A. x = Some y}" +"opt A == insert None {x. \y\A. x = Some y}" definition sup :: "'a ebinop \ 'a option ebinop" where "sup f o1 o2 == @@ -63,7 +63,7 @@ done lemma Some_le [iff]: - "(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)" + "(Some x <=_(le r) ox) = (\y. ox = Some y \ x <=_r y)" apply (unfold lesub_def le_def) apply (simp split: option.split) done @@ -89,11 +89,11 @@ lemma None_in_opt [iff]: - "None : opt A" + "None \ opt A" by (simp add: opt_def) lemma Some_in_opt [iff]: - "(Some x : opt A) = (x:A)" + "(Some x \ opt A) = (x\A)" apply (unfold opt_def) apply auto done @@ -132,34 +132,34 @@ have "closed ?A ?f" proof (unfold closed_def, intro strip) fix x y - assume x: "x : ?A" - assume y: "y : ?A" + assume x: "x \ ?A" + assume y: "y \ ?A" { fix a b assume ab: "x = OK a" "y = OK b" with x - have a: "\c. a = Some c \ c : A" + have a: "\c. a = Some c \ c \ A" by (clarsimp simp add: opt_def) from ab y - have b: "\d. b = Some d \ d : A" + have b: "\d. b = Some d \ d \ A" by (clarsimp simp add: opt_def) { fix c d assume "a = Some c" "b = Some d" with ab x y - have "c:A & d:A" + have "c \ A \ d \ A" by (simp add: err_def opt_def Bex_def) with clo - have "f c d : err A" + have "f c d \ err A" by (simp add: closed_def plussub_def err_def lift2_def) moreover fix z assume "f c d = OK z" ultimately - have "z : A" by simp + have "z \ A" by simp } note f_closed = this - have "sup f a b : ?A" + have "sup f a b \ ?A" proof (cases a) case None thus ?thesis @@ -171,7 +171,7 @@ qed } - thus "x +_?f y : ?A" + thus "x +_?f y \ ?A" by (simp add: plussub_def lift2_def split: err.split) qed @@ -209,7 +209,7 @@ have "\x\?A. \y\?A. \z\?A. x <=_?r z \ y <=_?r z \ x +_?f y <=_?r z" proof (intro strip, elim conjE) fix x y z - assume xyz: "x : ?A" "y : ?A" "z : ?A" + assume xyz: "x \ ?A" "y \ ?A" "z \ ?A" assume xz: "x <=_?r z" assume yz: "y <=_?r z" @@ -220,7 +220,7 @@ assume some: "a = Some d" "b = Some e" "c = Some g" with ok xyz - obtain "OK d:err A" "OK e:err A" "OK g:err A" + obtain "OK d \ err A" "OK e \ err A" "OK g \ err A" by simp with lub have "\ (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \ @@ -272,8 +272,8 @@ apply (unfold acc_def lesub_def le_def lesssub_def) apply (simp add: wf_eq_minimal split: option.split) apply clarify -apply (case_tac "? a. Some a : Q") - apply (erule_tac x = "{a . Some a : Q}" in allE) +apply (case_tac "\a. Some a \ Q") + apply (erule_tac x = "{a. Some a \ Q}" in allE) apply blast apply (case_tac "x") apply blast @@ -281,8 +281,8 @@ done lemma option_map_in_optionI: - "\ ox : opt S; !x:S. ox = Some x \ f x : S \ - \ map_option f ox : opt S" + "\ ox \ opt S; \x\S. ox = Some x \ f x \ S \ + \ map_option f ox \ opt S" apply (unfold map_option_case) apply (simp split: option.split) apply blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Product.thy --- a/src/HOL/MicroJava/DFA/Product.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Product.thy Thu Feb 15 13:04:36 2018 +0100 @@ -69,12 +69,12 @@ lemma plus_eq_Err_conv [simp]: - assumes "x:A" and "y:A" + assumes "x \ A" and "y \ A" and "semilat(err A, Err.le r, lift2 f)" - shows "(x +_f y = Err) = (~(? z:A. x <=_r z & y <=_r z))" + shows "(x +_f y = Err) = (\(\z\A. x <=_r z & y <=_r z))" proof - have plus_le_conv2: - "\r f z. \ z : err A; semilat (err A, r, f); OK x : err A; OK y : err A; + "\r f z. \ z \ err A; semilat (err A, r, f); OK x \ err A; OK y \ err A; OK x +_f OK y <=_r z\ \ OK x <=_r z \ OK y <=_r z" by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1]) from assms show ?thesis @@ -92,7 +92,7 @@ apply (case_tac "x +_f y") apply assumption apply (rename_tac "z") - apply (subgoal_tac "OK z: err A") + apply (subgoal_tac "OK z \ err A") apply (frule plus_le_conv2) apply assumption apply simp diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Semilat.thy --- a/src/HOL/MicroJava/DFA/Semilat.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Semilat.thy Thu Feb 15 13:04:36 2018 +0100 @@ -238,7 +238,7 @@ lemma is_lub_bigger1 [iff]: - "is_lub (r^* ) x y y = ((x,y)\r^* )" + "is_lub (r\<^sup>*) x y y = ((x,y)\r\<^sup>*)" (*<*) apply (unfold is_lub_def is_ub_def) apply blast @@ -246,7 +246,7 @@ (*>*) lemma is_lub_bigger2 [iff]: - "is_lub (r^* ) x y x = ((y,x)\r^* )" + "is_lub (r\<^sup>*) x y x = ((y,x)\r\<^sup>*)" (*<*) apply (unfold is_lub_def is_ub_def) apply blast @@ -254,12 +254,12 @@ (*>*) lemma extend_lub: - "\ single_valued r; is_lub (r^* ) x y u; (x',x) \ r \ - \ EX v. is_lub (r^* ) x' y v" + "\ single_valued r; is_lub (r\<^sup>*) x y u; (x',x) \ r \ + \ \v. is_lub (r\<^sup>*) x' y v" (*<*) apply (unfold is_lub_def is_ub_def) -apply (case_tac "(y,x) \ r^*") - apply (case_tac "(y,x') \ r^*") +apply (case_tac "(y,x) \ r\<^sup>*") + apply (case_tac "(y,x') \ r\<^sup>*") apply blast apply (blast elim: converse_rtranclE dest: single_valuedD) apply (rule exI) @@ -271,8 +271,8 @@ (*>*) lemma single_valued_has_lubs [rule_format]: - "\ single_valued r; (x,u) \ r^* \ \ (\y. (y,u) \ r^* \ - (EX z. is_lub (r^* ) x y z))" + "\single_valued r; (x,u) \ r\<^sup>*\ \ (\y. (y,u) \ r\<^sup>* \ + (\z. is_lub (r\<^sup>*) x y z))" (*<*) apply (erule converse_rtrancl_induct) apply clarify @@ -284,7 +284,7 @@ (*>*) lemma some_lub_conv: - "\ acyclic r; is_lub (r^* ) x y u \ \ some_lub (r^* ) x y = u" + "\acyclic r; is_lub (r\<^sup>*) x y u\ \ some_lub (r\<^sup>*) x y = u" (*<*) apply (unfold some_lub_def is_lub_def) apply (rule someI2) @@ -294,8 +294,8 @@ (*>*) lemma is_lub_some_lub: - "\ single_valued r; acyclic r; (x,u)\r^*; (y,u)\r^* \ - \ is_lub (r^* ) x y (some_lub (r^* ) x y)" + "\single_valued r; acyclic r; (x,u)\r\<^sup>*; (y,u)\r\<^sup>*\ + \ is_lub (r\<^sup>*) x y (some_lub (r\<^sup>*) x y)" (*<*) by (fastforce dest: single_valued_has_lubs simp add: some_lub_conv) (*>*) subsection\An executable lub-finder\ @@ -331,7 +331,7 @@ (*<*) apply(unfold exec_lub_def) apply(rule_tac P = "\z. (y,z) \ r\<^sup>* \ (z,u) \ r\<^sup>*" and - r = "(r \ {(a,b). (y,a) \ r\<^sup>* \ (b,u) \ r\<^sup>*})^-1" in while_rule) + r = "(r \ {(a,b). (y,a) \ r\<^sup>* \ (b,u) \ r\<^sup>*})\" in while_rule) apply(blast dest: is_lubD is_ubD) apply(erule conjE) apply(erule_tac z = u in converse_rtranclE) @@ -363,8 +363,8 @@ (*>*) lemma is_lub_exec_lub: - "\ single_valued r; acyclic r; (x,u):r^*; (y,u):r^*; \x y. (x,y) \ r \ f x = y \ - \ is_lub (r^* ) x y (exec_lub r f x y)" + "\ single_valued r; acyclic r; (x,u)\r\<^sup>*; (y,u)\r\<^sup>*; \x y. (x,y) \ r \ f x = y \ + \ is_lub (r\<^sup>*) x y (exec_lub r f x y)" (*<*) by (fastforce dest: single_valued_has_lubs simp add: exec_lub_conv) (*>*) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/SemilatAlg.thy --- a/src/HOL/MicroJava/DFA/SemilatAlg.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Thu Feb 15 13:04:36 2018 +0100 @@ -18,7 +18,7 @@ | "(x#xs) ++_f y = xs ++_f (x +_f y)" definition bounded :: "'s step_type \ nat \ bool" where -"bounded step n == !pps. \(q,t)\set(step p s). q nat \ 's set \ bool" where "pres_type step n A == \s\A. \p(q,s')\set (step p s). s' \ A" @@ -36,7 +36,7 @@ by (unfold mono_def, blast) lemma boundedD: - "\ bounded step n; p < n; (q,t) : set (step p xs) \ \ q < n" + "\ bounded step n; p < n; (q,t) \ set (step p xs) \ \ q < n" by (unfold bounded_def, blast) lemma lesubstep_type_refl [simp, intro]: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/DFA/Typing_Framework.thy --- a/src/HOL/MicroJava/DFA/Typing_Framework.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Typing_Framework.thy Thu Feb 15 13:04:36 2018 +0100 @@ -15,7 +15,7 @@ type_synonym 's step_type = "nat \ 's \ (nat \ 's) list" definition stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" where -"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q" +"stable r step ss p == \(q,s')\set(step p (ss!p)). s' <=_r ss!q" definition stables :: "'s ord \ 's step_type \ 's list \ bool" where "stables r step ss == !p 's \ 's step_type \ nat \ 's set \ ('s list \ 's list) \ bool" where -"is_bcv r T step n A bcv == !ss : list n A. - (!pss \ list n A. + (\pts \ list n A. ss <=[r] ts & wt_step r T step ts)" end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/J/Example.thy --- a/src/HOL/MicroJava/J/Example.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/J/Example.thy Thu Feb 15 13:04:36 2018 +0100 @@ -173,7 +173,7 @@ apply (simp (no_asm)) done -lemma not_Object_subcls [elim!]: "(Object, C) \ (subcls1 tprg)^+ ==> R" +lemma not_Object_subcls [elim!]: "(Object, C) \ (subcls1 tprg)\<^sup>+ ==> R" apply (auto dest!: tranclD subcls1D) done @@ -184,7 +184,7 @@ apply auto done -lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \ (subcls1 tprg)^+ ==> R" +lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \ (subcls1 tprg)\<^sup>+ ==> R" apply (auto dest!: tranclD subcls1D) done @@ -194,7 +194,7 @@ apply (auto split: if_split_asm simp add: map_of_Cons) done -lemma not_class_subcls_class [elim!]: "(C, C) \ (subcls1 tprg)^+ ==> R" +lemma not_class_subcls_class [elim!]: "(C, C) \ (subcls1 tprg)\<^sup>+ ==> R" apply (auto dest!: tranclD subcls1D) apply (frule class_tprgD) apply (auto dest!:) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/J/JBasis.thy --- a/src/HOL/MicroJava/J/JBasis.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/J/JBasis.thy Thu Feb 15 13:04:36 2018 +0100 @@ -21,17 +21,17 @@ "unique == distinct \ map fst" -lemma fst_in_set_lemma: "(x, y) : set xys ==> x : fst ` set xys" +lemma fst_in_set_lemma: "(x, y) \ set xys \ x \ fst ` set xys" by (induct xys) auto lemma unique_Nil [simp]: "unique []" by (simp add: unique_def) -lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))" +lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (\y. (x,y) \ set l))" by (auto simp: unique_def dest: fst_in_set_lemma) -lemma unique_append: "unique l' ==> unique l ==> - (!(x,y):set l. !(x',y'):set l'. x' ~= x) ==> unique (l @ l')" +lemma unique_append: "unique l' \ unique l \ + (\(x,y) \ set l. \(x',y') \ set l'. x' \ x) \ unique (l @ l')" by (induct l) (auto dest: fst_in_set_lemma) lemma unique_map_inj: "unique l ==> inj f ==> unique (map (%(k,x). (f k, g k x)) l)" @@ -40,7 +40,7 @@ subsection "More about Maps" -lemma map_of_SomeI: "unique l ==> (k, x) : set l ==> map_of l k = Some x" +lemma map_of_SomeI: "unique l \ (k, x) \ set l \ map_of l k = Some x" by (induct l) auto lemma Ball_set_table: "(\(x,y)\set l. P x y) ==> (\x. \y. map_of l x = Some y --> P x y)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/J/TypeRel.thy --- a/src/HOL/MicroJava/J/TypeRel.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/J/TypeRel.thy Thu Feb 15 13:04:36 2018 +0100 @@ -20,7 +20,7 @@ abbreviation subcls :: "'c prog => cname \ cname => bool" ("_ \ _ \C _" [71,71,71] 70) - where "G \ C \C D \ (C, D) \ (subcls1 G)^*" + where "G \ C \C D \ (C, D) \ (subcls1 G)\<^sup>*" lemma subcls1D: "G\C\C1D \ C \ Object \ (\fs ms. class G C = Some (D,fs,ms))" @@ -40,7 +40,7 @@ apply auto done -lemma subcls_is_class: "(C, D) \ (subcls1 G)^+ ==> is_class G C" +lemma subcls_is_class: "(C, D) \ (subcls1 G)\<^sup>+ \ is_class G C" apply (unfold is_class_def) apply(erule trancl_trans_induct) apply (auto dest!: subcls1D) @@ -56,21 +56,21 @@ definition class_rec :: "'c prog \ cname \ 'a \ (cname \ fdecl list \ 'c mdecl list \ 'a \ 'a) \ 'a" where - "class_rec G == wfrec ((subcls1 G)^-1) + "class_rec G == wfrec ((subcls1 G)\) (\r C t f. case class G C of None \ undefined | Some (D,fs,ms) \ f C fs ms (if C = Object then t else r D t f))" lemma class_rec_lemma: - assumes wf: "wf ((subcls1 G)^-1)" + assumes wf: "wf ((subcls1 G)\)" and cls: "class G C = Some (D, fs, ms)" shows "class_rec G C t f = f C fs ms (if C=Object then t else class_rec G D t f)" by (subst wfrec_def_adm[OF class_rec_def]) (auto simp: assms adm_wf_def fun_eq_iff subcls1I split: option.split) definition - "wf_class G = wf ((subcls1 G)^-1)" + "wf_class G = wf ((subcls1 G)\)" @@ -89,7 +89,7 @@ subcls1 . -definition subcls' where "subcls' G = (subcls1p G)^**" +definition subcls' where "subcls' G = (subcls1p G)\<^sup>*\<^sup>*" code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) @@ -98,7 +98,7 @@ . lemma subcls_conv_subcls' [code_unfold]: - "(subcls1 G)^* = {(C, D). subcls' G C D}" + "(subcls1 G)\<^sup>* = {(C, D). subcls' G C D}" by(simp add: subcls'_def subcls1_def rtrancl_def) lemma class_rec_code [code]: @@ -126,8 +126,8 @@ "wf_class G \ (\(C, rest) \ set G. C \ Object \ \ G \ fst (the (class G C)) \C C)" proof assume "wf_class G" - hence wf: "wf (((subcls1 G)^+)^-1)" unfolding wf_class_def by(rule wf_converse_trancl) - hence acyc: "acyclic ((subcls1 G)^+)" by(auto dest: wf_acyclic) + hence wf: "wf (((subcls1 G)\<^sup>+)\)" unfolding wf_class_def by(rule wf_converse_trancl) + hence acyc: "acyclic ((subcls1 G)\<^sup>+)" by(auto dest: wf_acyclic) show "\(C, rest) \ set G. C \ Object \ \ G \ fst (the (class G C)) \C C" proof(safe) fix C D fs ms @@ -138,9 +138,9 @@ where "class": "class G C = Some (D', fs', ms')" unfolding class_def by(auto dest!: weak_map_of_SomeI) hence "G \ C \C1 D'" using \C \ Object\ .. - hence *: "(C, D') \ (subcls1 G)^+" .. + hence *: "(C, D') \ (subcls1 G)\<^sup>+" .. also from * acyc have "C \ D'" by(auto simp add: acyclic_def) - with subcls "class" have "(D', C) \ (subcls1 G)^+" by(auto dest: rtranclD) + with subcls "class" have "(D', C) \ (subcls1 G)\<^sup>+" by(auto dest: rtranclD) finally show False using acyc by(auto simp add: acyclic_def) qed next @@ -189,7 +189,7 @@ definition field :: "'c prog \ cname => (vname \ cname \ ty)" where [code]: "field == map_of o (map (\((fn,fd),ft). (fn,(fd,ft)))) o fields" -lemma method_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==> +lemma method_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)\)|] ==> method (G,C) = (if C = Object then empty else method (G,D)) ++ map_of (map (\(s,m). (s,(C,m))) ms)" apply (unfold method_def) @@ -198,7 +198,7 @@ apply auto done -lemma fields_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==> +lemma fields_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)\)|] ==> fields (G,C) = map (\(fn,ft). ((fn,C),ft)) fs @ (if C = Object then [] else fields (G,D))" apply (unfold fields_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/J/WellForm.thy --- a/src/HOL/MicroJava/J/WellForm.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/J/WellForm.thy Thu Feb 15 13:04:36 2018 +0100 @@ -134,7 +134,7 @@ apply (auto intro!: map_of_SomeI) done -lemma subcls1_wfD: "[|G\C\C1D; ws_prog G|] ==> D \ C \ (D, C) \ (subcls1 G)^+" +lemma subcls1_wfD: "[|G\C\C1D; ws_prog G|] ==> D \ C \ (D, C) \ (subcls1 G)\<^sup>+" apply( frule trancl.r_into_trancl [where r="subcls1 G"]) apply( drule subcls1D) apply(clarify) @@ -149,13 +149,13 @@ apply (auto split: option.split_asm) done -lemma subcls_asym: "[|ws_prog G; (C, D) \ (subcls1 G)^+|] ==> (D, C) \ (subcls1 G)^+" +lemma subcls_asym: "[|ws_prog G; (C, D) \ (subcls1 G)\<^sup>+|] ==> (D, C) \ (subcls1 G)\<^sup>+" apply(erule trancl.cases) apply(fast dest!: subcls1_wfD ) apply(fast dest!: subcls1_wfD intro: trancl_trans) done -lemma subcls_irrefl: "[|ws_prog G; (C, D) \ (subcls1 G)^+|] ==> C \ D" +lemma subcls_irrefl: "[|ws_prog G; (C, D) \ (subcls1 G)\<^sup>+|] ==> C \ D" apply (erule trancl_trans_induct) apply (auto dest: subcls1_wfD subcls_asym) done @@ -165,7 +165,7 @@ apply (fast dest: subcls_irrefl) done -lemma wf_subcls1: "ws_prog G ==> wf ((subcls1 G)^-1)" +lemma wf_subcls1: "ws_prog G ==> wf ((subcls1 G)\)" apply (rule finite_acyclic_wf) apply ( subst finite_converse) apply ( rule finite_subcls1) @@ -174,7 +174,7 @@ done lemma subcls_induct_struct: -"[|ws_prog G; !!C. \D. (C, D) \ (subcls1 G)^+ --> P D ==> P C|] ==> P C" +"[|ws_prog G; !!C. \D. (C, D) \ (subcls1 G)\<^sup>+ --> P D ==> P C|] ==> P C" (is "?A \ PROP ?P \ _") proof - assume p: "PROP ?P" @@ -189,7 +189,7 @@ qed lemma subcls_induct: -"[|wf_prog wf_mb G; !!C. \D. (C, D) \ (subcls1 G)^+ --> P D ==> P C|] ==> P C" +"[|wf_prog wf_mb G; !!C. \D. (C, D) \ (subcls1 G)\<^sup>+ --> P D ==> P C|] ==> P C" (is "?A \ PROP ?P \ _") by (fact subcls_induct_struct [OF wf_prog_ws_prog]) @@ -367,7 +367,7 @@ done lemma fields_mono_lemma [rule_format (no_asm)]: - "[|ws_prog G; (C', C) \ (subcls1 G)^*|] ==> + "[|ws_prog G; (C', C) \ (subcls1 G)\<^sup>*|] ==> x \ set (fields (G,C)) --> x \ set (fields (G,C'))" apply(erule converse_rtrancl_induct) apply( safe dest!: subcls1D) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/MicroJava/JVM/JVMExec.thy --- a/src/HOL/MicroJava/JVM/JVMExec.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/MicroJava/JVM/JVMExec.thy Thu Feb 15 13:04:36 2018 +0100 @@ -25,7 +25,7 @@ definition exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool" ("_ \ _ \jvm\ _" [61,61,61]60) where - "G \ s \jvm\ t == (s,t) \ {(s,t). exec(G,s) = Some t}^*" + "G \ s \jvm\ t == (s,t) \ {(s,t). exec(G,s) = Some t}\<^sup>*" text \ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/NanoJava/Equivalence.thy --- a/src/HOL/NanoJava/Equivalence.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/NanoJava/Equivalence.thy Thu Feb 15 13:04:36 2018 +0100 @@ -173,7 +173,7 @@ lemma MGF_Loop: "\Z. A \ {(=) Z} c {\t. \n. Z -c-n\ t} \ A \ {(=) Z} While (x) c {\t. \n. Z -While (x) c-n\ t}" -apply (rule_tac P' = "\Z s. (Z,s) \ ({(s,t). \n. s \ Null \ s -c-n\ t})^*" +apply (rule_tac P' = "\Z s. (Z,s) \ ({(s,t). \n. s \ Null \ s -c-n\ t})\<^sup>*" in hoare_ehoare.Conseq) apply (rule allI) apply (rule hoare_ehoare.Loop) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/NanoJava/TypeRel.thy Thu Feb 15 13:04:36 2018 +0100 @@ -19,7 +19,7 @@ where "C \C1 D == (C,D) \ subcls1" abbreviation subcls_syntax :: "[cname, cname] => bool" ("_ \C _" [71,71] 70) - where "C \C D == (C,D) \ subcls1^*" + where "C \C D \ (C,D) \ subcls1\<^sup>*" subsection "Declarations and properties not used in the meta theory" @@ -59,20 +59,20 @@ definition ws_prog :: "bool" where "ws_prog \ \(C,c)\set Prog. C\Object \ - is_class (super c) \ (super c,C)\subcls1^+" + is_class (super c) \ (super c,C)\subcls1\<^sup>+" lemma ws_progD: "\class C = Some c; C\Object; ws_prog\ \ - is_class (super c) \ (super c,C)\subcls1^+" + is_class (super c) \ (super c,C)\subcls1\<^sup>+" apply (unfold ws_prog_def class_def) apply (drule_tac map_of_SomeD) apply auto done -lemma subcls1_irrefl_lemma1: "ws_prog \ subcls1^-1 \ subcls1^+ = {}" +lemma subcls1_irrefl_lemma1: "ws_prog \ subcls1\ \ subcls1\<^sup>+ = {}" by (fast dest: subcls1D ws_progD) (* irrefl_tranclI in Transitive_Closure.thy is more general *) -lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+" +lemma irrefl_tranclI': "r\ \ r\<^sup>+ = {} \ \x. (x, x) \ r\<^sup>+" by(blast elim: tranclE dest: trancl_into_rtrancl) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nitpick_Examples/Core_Nits.thy --- a/src/HOL/Nitpick_Examples/Core_Nits.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nitpick_Examples/Core_Nits.thy Thu Feb 15 13:04:36 2018 +0100 @@ -55,7 +55,7 @@ nitpick [card = 1, expect = genuine] oops -lemma "{(a::'a\'a, b::'b)}^-1 = {(b, a)}" +lemma "{(a::'a\'a, b::'b)}\ = {(b, a)}" nitpick [card = 1-12, expect = none] by auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nitpick_Examples/Mini_Nits.thy --- a/src/HOL/Nitpick_Examples/Mini_Nits.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Thu Feb 15 13:04:36 2018 +0100 @@ -90,7 +90,7 @@ ML \genuine 2 @{prop "fst (a, b) = b"}\ ML \genuine 2 @{prop "fst (a, b) \ b"}\ ML \genuine 2 @{prop "f ((x, z), y) = (x, z)"}\ -ML \none 2 @{prop "(ALL x. f x = fst x) \ f ((x, z), y) = (x, z)"}\ +ML \none 2 @{prop "(\x. f x = fst x) \ f ((x, z), y) = (x, z)"}\ ML \none 4 @{prop "snd (a, b) = b"}\ ML \none 1 @{prop "snd (a, b) = a"}\ ML \genuine 2 @{prop "snd (a, b) = a"}\ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nitpick_Examples/Refute_Nits.thy --- a/src/HOL/Nitpick_Examples/Refute_Nits.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy Thu Feb 15 13:04:36 2018 +0100 @@ -245,15 +245,15 @@ text \``The transitive closure of an arbitrary relation is non-empty.''\ definition "trans" :: "('a \ 'a \ bool) \ bool" where -"trans P \ (ALL x y z. P x y \ P y z \ P x z)" +"trans P \ (\x y z. P x y \ P y z \ P x z)" definition "subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where -"subset P Q \ (ALL x y. P x y \ Q x y)" +"subset P Q \ (\x y. P x y \ Q x y)" definition "trans_closure" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where -"trans_closure P Q \ (subset Q P) \ (trans P) \ (ALL R. subset Q R \ trans R \ subset P R)" +"trans_closure P Q \ (subset Q P) \ (trans P) \ (\R. subset Q R \ trans R \ subset P R)" lemma "trans_closure T P \ (\x y. T x y)" nitpick [expect = genuine] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nominal/Examples/Class1.thy --- a/src/HOL/Nominal/Examples/Class1.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nominal/Examples/Class1.thy Thu Feb 15 13:04:36 2018 +0100 @@ -5165,7 +5165,7 @@ abbreviation a_star_redu :: "trm \ trm \ bool" ("_ \\<^sub>a* _" [100,100] 100) where - "M \\<^sub>a* M' \ (a_redu)^** M M'" + "M \\<^sub>a* M' \ (a_redu)\<^sup>*\<^sup>* M M'" lemma a_starI: assumes a: "M \\<^sub>a M'" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nonstandard_Analysis/HyperDef.thy --- a/src/HOL/Nonstandard_Analysis/HyperDef.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nonstandard_Analysis/HyperDef.thy Thu Feb 15 13:04:36 2018 +0100 @@ -138,7 +138,7 @@ lemma lemma_starrel_refl [simp]: "x \ starrel `` {x}" by (simp add: starrel_def) -lemma starrel_in_hypreal [simp]: "starrel``{x}:star" +lemma starrel_in_hypreal [simp]: "starrel``{x}\star" by (simp add: star_def starrel_def quotient_def, blast) declare Abs_star_inject [simp] Abs_star_inverse [simp] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nonstandard_Analysis/NSCA.thy --- a/src/HOL/Nonstandard_Analysis/NSCA.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nonstandard_Analysis/NSCA.thy Thu Feb 15 13:04:36 2018 +0100 @@ -57,13 +57,13 @@ by (auto simp add: Standard_def image_def) (metis inj_star_of inv_f_f) lemma SComplex_hcomplex_of_complex_image: - "[| \x. x: P; P \ SComplex |] ==> \Q. P = hcomplex_of_complex ` Q" + "\\x. x \ P; P \ SComplex\ \ \Q. P = hcomplex_of_complex ` Q" apply (simp add: Standard_def, blast) done lemma SComplex_SReal_dense: - "[| x \ SComplex; y \ SComplex; hcmod x < hcmod y - |] ==> \r \ Reals. hcmod x< r \ r < hcmod y" + "\x \ SComplex; y \ SComplex; hcmod x < hcmod y + \ \ \r \ Reals. hcmod x< r \ r < hcmod y" apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex) done @@ -78,7 +78,7 @@ by (simp add: HFinite_def) lemma HFinite_bounded_hcmod: - "[|x \ HFinite; y \ hcmod x; 0 \ y |] ==> y: HFinite" + "\x \ HFinite; y \ hcmod x; 0 \ y\ \ y \ HFinite" by (auto intro: HFinite_bounded simp add: HFinite_hcmod_iff) @@ -343,14 +343,14 @@ by (simp add: HComplex_def) (* Here we go - easy proof now!! *) -lemma stc_part_Ex: "x:HFinite ==> \t \ SComplex. x \ t" +lemma stc_part_Ex: "x \ HFinite \ \t \ SComplex. x \ t" apply (simp add: hcomplex_HFinite_iff hcomplex_approx_iff) apply (rule_tac x="HComplex (st (hRe x)) (st (hIm x))" in bexI) apply (simp add: st_approx_self [THEN approx_sym]) apply (simp add: Standard_HComplex st_SReal [unfolded Reals_eq_Standard]) done -lemma stc_part_Ex1: "x\HFinite \ \!t. t \ SComplex \ x \ t" +lemma stc_part_Ex1: "x \ HFinite \ \!t. t \ SComplex \ x \ t" apply (drule stc_part_Ex, safe) apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym) apply (auto intro!: approx_unique_complex) @@ -362,7 +362,7 @@ subsection\Theorems About Monads\ -lemma monad_zero_hcmod_iff: "(x \ monad 0) = (hcmod x:monad 0)" +lemma monad_zero_hcmod_iff: "(x \ monad 0) = (hcmod x \ monad 0)" by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Nonstandard_Analysis/Star.thy --- a/src/HOL/Nonstandard_Analysis/Star.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Nonstandard_Analysis/Star.thy Thu Feb 15 13:04:36 2018 +0100 @@ -231,7 +231,7 @@ by transfer (rule refl) text \General lemma/theorem needed for proofs in elementary topology of the reals.\ -lemma starfun_mem_starset: "\x. ( *f* f) x : *s* A \ x \ *s* {x. f x \ A}" +lemma starfun_mem_starset: "\x. ( *f* f) x \ *s* A \ x \ *s* {x. f x \ A}" by transfer simp text \Alternative definition for \hrabs\ with \rabs\ function applied diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Predicate_Compile.thy --- a/src/HOL/Predicate_Compile.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Predicate_Compile.thy Thu Feb 15 13:04:36 2018 +0100 @@ -32,10 +32,10 @@ \ definition contains :: "'a set => 'a => bool" -where "contains A x \ x : A" +where "contains A x \ x \ A" definition contains_pred :: "'a set => 'a => unit Predicate.pred" -where "contains_pred A x = (if x : A then Predicate.single () else bot)" +where "contains_pred A x = (if x \ A then Predicate.single () else bot)" lemma pred_of_setE: assumes "Predicate.eval (pred_of_set A) x" @@ -52,7 +52,7 @@ by(simp add: contains_def) lemma containsE: assumes "contains A x" - obtains A' x' where "A = A'" "x = x'" "x : A" + obtains A' x' where "A = A'" "x = x'" "x \ A" using assms by(simp add: contains_def) lemma contains_predI: "contains A x ==> Predicate.eval (contains_pred A x) ()" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Predicate_Compile_Examples/Hotel_Example.thy --- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Thu Feb 15 13:04:36 2018 +0100 @@ -66,10 +66,10 @@ primrec hotel :: "event list \ bool" where "hotel [] = True" -| "hotel (e # s) = (hotel s & (case e of +| "hotel (e # s) = (hotel s \ (case e of Check_in g r (k,k') \ k = currk s r \ k' \ issued s | - Enter g r (k,k') \ (k,k') : cards s g & (roomk s r : {k, k'}) | - Exit g r \ g : isin s r))" + Enter g r (k,k') \ (k,k') \ cards s g \ (roomk s r \ {k, k'}) | + Exit g r \ g \ isin s r))" definition no_Check_in :: "event list \ room \ bool" where(*>*) [code del]: "no_Check_in s r \ \(\g c. Check_in g r c \ set s)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy --- a/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Thu Feb 15 13:04:36 2018 +0100 @@ -59,11 +59,11 @@ declare List.member_rec[code_pred_def] -lemma [code_pred_def]: "no_Check_in s r = (~ (EX g c. List.member s (Check_in g r c)))" +lemma [code_pred_def]: "no_Check_in s r = (\ (\g c. List.member s (Check_in g r c)))" unfolding no_Check_in_def member_def by auto lemma [code_pred_def]: "feels_safe s r = -(EX s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'. +(\s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'. s = s\<^sub>3 @ [Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 & diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Presburger.thy Thu Feb 15 13:04:36 2018 +0100 @@ -194,12 +194,12 @@ hence "P 0" using P Pmod by simp moreover have "P 0 = P(0 - (-1)*d)" using modd by blast ultimately have "P d" by simp - moreover have "d : {1..d}" using dpos by simp + moreover have "d \ {1..d}" using dpos by simp ultimately show ?RHS .. next assume not0: "x mod d \ 0" have "P(x mod d)" using dpos P Pmod by simp - moreover have "x mod d : {1..d}" + moreover have "x mod d \ {1..d}" proof - from dpos have "0 \ x mod d" by(rule pos_mod_sign) moreover from dpos have "x mod d < d" by(rule pos_mod_bound) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Prolog/Test.thy --- a/src/HOL/Prolog/Test.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Prolog/Test.thy Thu Feb 15 13:04:36 2018 +0100 @@ -52,9 +52,9 @@ where append: "\x xs ys zs. append [] xs xs .. append (x#xs) ys (x#zs) :- append xs ys zs" and - reverse: "\L1 L2. reverse L1 L2 :- (!rev_aux. - (!L. rev_aux [] L L ).. - (!X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3)) + reverse: "\L1 L2. reverse L1 L2 :- (\rev_aux. + (\L. rev_aux [] L L ).. + (\X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3)) => rev_aux L1 L2 [])" and mappred: "\x xs y ys P. mappred P [] [] .. @@ -70,7 +70,7 @@ (* actual definitions of empty and add is hidden -> yields abstract data type *) - bag_appl: "\A B X Y. bag_appl A B X Y:- (? S1 S2 S3 S4 S5. + bag_appl: "\A B X Y. bag_appl A B X Y:- (\S1 S2 S3 S4 S5. empty S1 & add A S1 S2 & add B S2 S3 & @@ -92,7 +92,7 @@ apply (prolog prog_Test) done -schematic_goal "!y. append [a,b] y (?L y)" +schematic_goal "\y. append [a,b] y (?L y)" apply (prolog prog_Test) done @@ -117,7 +117,7 @@ back done -schematic_goal "mappred (%x y. ? z. age z y) ?x [23,24]" +schematic_goal "mappred (\x y. \z. age z y) ?x [23,24]" apply (prolog prog_Test) done @@ -159,11 +159,11 @@ back done -lemma "? x y. age x y" +lemma "\x y. age x y" apply (prolog prog_Test) done -lemma "!x y. append [] x x" +lemma "\x y. append [] x x" apply (prolog prog_Test) done @@ -181,18 +181,18 @@ done (*reset trace_DEPTH_FIRST;*) -schematic_goal "(!x. age x 25 :- age x 23) => age ?x 25 & age ?y 25" +schematic_goal "(\x. age x 25 :- age x 23) => age ?x 25 & age ?y 25" apply (prolog prog_Test) back back back done -lemma "!x. ? y. eq x y" +lemma "\x. \y. eq x y" apply (prolog prog_Test) done -schematic_goal "? P. P & eq P ?x" +schematic_goal "\P. P \ eq P ?x" apply (prolog prog_Test) (* back @@ -206,20 +206,20 @@ *) done -lemma "? P. eq P (True & True) & P" +lemma "\P. eq P (True & True) \ P" apply (prolog prog_Test) done -lemma "? P. eq P (|) & P k True" +lemma "\P. eq P (\) \ P k True" apply (prolog prog_Test) done -lemma "? P. eq P (Q => True) & P" +lemma "\P. eq P (Q => True) \ P" apply (prolog prog_Test) done (* P flexible: *) -lemma "(!P k l. P k l :- eq P Q) => Q a b" +lemma "(\P k l. P k l :- eq P Q) => Q a b" apply (prolog prog_Test) done (* @@ -228,11 +228,11 @@ *) (* implication and disjunction in atom: *) -lemma "? Q. (!p q. R(q :- p) => R(Q p q)) & Q (t | s) (s | t)" +lemma "\Q. (\p q. R(q :- p) => R(Q p q)) \ Q (t | s) (s | t)" by fast (* disjunction in atom: *) -lemma "(!P. g P :- (P => b | a)) => g(a | b)" +lemma "(\P. g P :- (P => b \ a)) => g(a \ b)" apply (tactic "step_tac (put_claset HOL_cs @{context}) 1") apply (tactic "step_tac (put_claset HOL_cs @{context}) 1") apply (tactic "step_tac (put_claset HOL_cs @{context}) 1") @@ -247,21 +247,21 @@ by fast *) -schematic_goal "!Emp Stk.( +schematic_goal "\Emp Stk.( empty (Emp::'b) .. - (!(X::nat) S. add X (S::'b) (Stk X S)) .. - (!(X::nat) S. remove X ((Stk X S)::'b) S)) + (\(X::nat) S. add X (S::'b) (Stk X S)) .. + (\(X::nat) S. remove X ((Stk X S)::'b) S)) => bag_appl 23 24 ?X ?Y" oops -schematic_goal "!Qu. ( - (!L. empty (Qu L L)) .. - (!(X::nat) L K. add X (Qu L (X#K)) (Qu L K)) .. - (!(X::nat) L K. remove X (Qu (X#L) K) (Qu L K))) +schematic_goal "\Qu. ( + (\L. empty (Qu L L)) .. + (\(X::nat) L K. add X (Qu L (X#K)) (Qu L K)) .. + (\(X::nat) L K. remove X (Qu (X#L) K) (Qu L K))) => bag_appl 23 24 ?X ?Y" oops -lemma "D & (!y. E) :- (!x. True & True) :- True => D" +lemma "D \ (\y. E) :- (\x. True \ True) :- True => D" apply (prolog prog_Test) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Prolog/Type.thy --- a/src/HOL/Prolog/Type.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Prolog/Type.thy Thu Feb 15 13:04:36 2018 +0100 @@ -13,14 +13,14 @@ axiomatization bool :: ty and nat :: ty and - arrow :: "ty => ty => ty" (infixr "->" 20) and - typeof :: "[tm, ty] => bool" and + arrow :: "ty \ ty \ ty" (infixr "->" 20) and + typeof :: "[tm, ty] \ bool" and anyterm :: tm where common_typeof: " -typeof (app M N) B :- typeof M (A -> B) & typeof N A.. +typeof (app M N) B :- typeof M (A -> B) \ typeof N A.. -typeof (cond C L R) A :- typeof C bool & typeof L A & typeof R A.. -typeof (fix F) A :- (!x. typeof x A => typeof (F x) A).. +typeof (cond C L R) A :- typeof C bool \ typeof L A \ typeof R A.. +typeof (fix F) A :- (\x. typeof x A => typeof (F x) A).. typeof true bool.. typeof false bool.. @@ -35,7 +35,7 @@ typeof (M * N) nat :- typeof M nat & typeof N nat" axiomatization where good_typeof: " -typeof (abs Bo) (A -> B) :- (!x. typeof x A => typeof (Bo x) B)" +typeof (abs Bo) (A -> B) :- (\x. typeof x A => typeof (Bo x) B)" axiomatization where bad1_typeof: " typeof (abs Bo) (A -> B) :- (typeof varterm A => typeof (Bo varterm) B)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/Commutation.thy --- a/src/HOL/Proofs/Lambda/Commutation.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/Commutation.thy Thu Feb 15 13:04:36 2018 +0100 @@ -30,11 +30,11 @@ definition Church_Rosser :: "('a => 'a => bool) => bool" where "Church_Rosser R = - (\x y. (sup R (R^--1))^** x y --> (\z. R^** x z \ R^** y z))" + (\x y. (sup R (R\\))\<^sup>*\<^sup>* x y \ (\z. R\<^sup>*\<^sup>* x z \ R\<^sup>*\<^sup>* y z))" abbreviation confluent :: "('a => 'a => bool) => bool" where - "confluent R == diamond (R^**)" + "confluent R == diamond (R\<^sup>*\<^sup>*)" subsection \Basic lemmas\ @@ -53,13 +53,13 @@ done lemma square_reflcl: - "[| square R S T (R^==); S \ T |] ==> square (R^==) S T (R^==)" + "[| square R S T (R\<^sup>=\<^sup>=); S \ T |] ==> square (R\<^sup>=\<^sup>=) S T (R\<^sup>=\<^sup>=)" apply (unfold square_def) apply (blast dest: predicate2D) done lemma square_rtrancl: - "square R S S T ==> square (R^**) S S (T^**)" + "square R S S T ==> square (R\<^sup>*\<^sup>*) S S (T\<^sup>*\<^sup>*)" apply (unfold square_def) apply (intro strip) apply (erule rtranclp_induct) @@ -68,7 +68,7 @@ done lemma square_rtrancl_reflcl_commute: - "square R S (S^**) (R^==) ==> commute (R^**) (S^**)" + "square R S (S\<^sup>*\<^sup>*) (R\<^sup>=\<^sup>=) ==> commute (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*)" apply (unfold commute_def) apply (fastforce dest: square_reflcl square_sym [THEN square_rtrancl]) done @@ -81,7 +81,7 @@ apply (blast intro: square_sym) done -lemma commute_rtrancl: "commute R S ==> commute (R^**) (S^**)" +lemma commute_rtrancl: "commute R S ==> commute (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*)" apply (unfold commute_def) apply (blast intro: square_rtrancl square_sym) done @@ -107,19 +107,19 @@ done lemma square_reflcl_confluent: - "square R R (R^==) (R^==) ==> confluent R" + "square R R (R\<^sup>=\<^sup>=) (R\<^sup>=\<^sup>=) ==> confluent R" apply (unfold diamond_def) apply (fast intro: square_rtrancl_reflcl_commute elim: square_subset) done lemma confluent_Un: - "[| confluent R; confluent S; commute (R^**) (S^**) |] ==> confluent (sup R S)" + "[| confluent R; confluent S; commute (R\<^sup>*\<^sup>*) (S\<^sup>*\<^sup>*) |] ==> confluent (sup R S)" apply (rule rtranclp_sup_rtranclp [THEN subst]) apply (blast dest: diamond_Un intro: diamond_confluent) done lemma diamond_to_confluence: - "[| diamond R; T \ R; R \ T^** |] ==> confluent T" + "[| diamond R; T \ R; R \ T\<^sup>*\<^sup>* |] ==> confluent T" apply (force intro: diamond_confluent dest: rtranclp_subset [symmetric]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/Eta.thy --- a/src/HOL/Proofs/Lambda/Eta.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/Eta.thy Thu Feb 15 13:04:36 2018 +0100 @@ -27,11 +27,11 @@ abbreviation eta_reds :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>*" 50) where - "s \\<^sub>\\<^sup>* t == eta^** s t" + "s \\<^sub>\\<^sup>* t == eta\<^sup>*\<^sup>* s t" abbreviation eta_red0 :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>=" 50) where - "s \\<^sub>\\<^sup>= t == eta^== s t" + "s \\<^sub>\\<^sup>= t == eta\<^sup>=\<^sup>= s t" inductive_cases eta_cases [elim!]: "Abs s \\<^sub>\ z" @@ -79,7 +79,7 @@ subsection \Confluence of \eta\\ -lemma square_eta: "square eta eta (eta^==) (eta^==)" +lemma square_eta: "square eta eta (eta\<^sup>=\<^sup>=) (eta\<^sup>=\<^sup>=)" apply (unfold square_def id_def) apply (rule impI [THEN allI [THEN allI]]) apply (erule eta.induct) @@ -148,7 +148,7 @@ shows "u[s/i] \\<^sub>\\<^sup>* u[t/i]" using eta by induct (iprover intro: rtrancl_eta_subst rtranclp_trans)+ -lemma square_beta_eta: "square beta eta (eta^**) (beta^==)" +lemma square_beta_eta: "square beta eta (eta\<^sup>*\<^sup>*) (beta\<^sup>=\<^sup>=)" apply (unfold square_def) apply (rule impI [THEN allI [THEN allI]]) apply (erule beta.induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/InductTermi.thy --- a/src/HOL/Proofs/Lambda/InductTermi.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/InductTermi.thy Thu Feb 15 13:04:36 2018 +0100 @@ -84,7 +84,7 @@ apply (erule mp) apply clarify apply (drule_tac r=beta in conversepI) - apply (drule_tac r="beta^--1" in ex_step1I, assumption) + apply (drule_tac r="beta\\" in ex_step1I, assumption) apply clarify apply (rename_tac us) apply (erule_tac x = "Var n \\ us" in allE) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/Lambda.thy --- a/src/HOL/Proofs/Lambda/Lambda.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/Lambda.thy Thu Feb 15 13:04:36 2018 +0100 @@ -65,7 +65,7 @@ abbreviation beta_reds :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>*" 50) where - "s \\<^sub>\\<^sup>* t == beta^** s t" + "s \\<^sub>\\<^sup>* t == beta\<^sup>*\<^sup>* s t" inductive_cases beta_cases [elim!]: "Var i \\<^sub>\ t" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/ListOrder.thy --- a/src/HOL/Proofs/Lambda/ListOrder.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/ListOrder.thy Thu Feb 15 13:04:36 2018 +0100 @@ -24,12 +24,12 @@ us @ z' # vs)" -lemma step1_converse [simp]: "step1 (r^--1) = (step1 r)^--1" +lemma step1_converse [simp]: "step1 (r\\) = (step1 r)\\" apply (unfold step1_def) apply (blast intro!: order_antisym) done -lemma in_step1_converse [iff]: "(step1 (r^--1) x y) = ((step1 r)^--1 x y)" +lemma in_step1_converse [iff]: "(step1 (r\\) x y) = ((step1 r)\\ x y)" apply auto done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Proofs/Lambda/ParRed.thy --- a/src/HOL/Proofs/Lambda/ParRed.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Proofs/Lambda/ParRed.thy Thu Feb 15 13:04:36 2018 +0100 @@ -30,7 +30,7 @@ subsection \Inclusions\ -text \\beta \ par_beta \ beta^*\ \medskip\ +text \\beta \ par_beta \ beta\<^sup>*\ \medskip\ lemma par_beta_varL [simp]: "(Var n => t) = (t = Var n)" @@ -45,7 +45,7 @@ apply (blast intro!: par_beta_refl)+ done -lemma par_beta_subset_beta: "par_beta <= beta^**" +lemma par_beta_subset_beta: "par_beta \ beta\<^sup>*\<^sup>*" apply (rule predicate2I) apply (erule par_beta.induct) apply blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Quickcheck_Examples/Hotel_Example.thy --- a/src/HOL/Quickcheck_Examples/Hotel_Example.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Quickcheck_Examples/Hotel_Example.thy Thu Feb 15 13:04:36 2018 +0100 @@ -70,8 +70,8 @@ "hotel [] = True" | "hotel (e # s) = (hotel s & (case e of Check_in g r (k,k') \ k = currk s r \ k' \ issued s | - Enter g r (k,k') \ (k,k') : cards s g & (roomk s r : {k, k'}) | - Exit g r \ g : isin s r))" + Enter g r (k,k') \ (k,k') \ cards s g & (roomk s r \ {k, k'}) | + Exit g r \ g \ isin s r))" definition no_Check_in :: "event list \ room \ bool" where(*>*) [code del]: "no_Check_in s r \ \(\g c. Check_in g r c \ set s)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -303,7 +303,7 @@ (* FIXME: some dramatic performance decrease after changing the code equation of the ntrancl *) lemma - "(x, z) : rtrancl (R Un S) ==> \ y. (x, y) : rtrancl R & (y, z) : rtrancl S" + "(x, z) \ rtrancl (R Un S) \ \y. (x, y) \ rtrancl R \ (y, z) \ rtrancl S" (*quickcheck[exhaustive, expect = counterexample]*) oops diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -41,34 +41,34 @@ oops lemma - "x > 2 --> (\y :: nat. x < y & y <= 2)" + "x > 2 \ (\y :: nat. x < y \ y \ 2)" quickcheck[tester = narrowing, finite_types = false, size = 10] oops lemma - "\ x. \ y :: nat. x > 3 --> (y < x & y > 3)" + "\x. \y :: nat. x > 3 \ (y < x \ y > 3)" quickcheck[tester = narrowing, finite_types = false, size = 7] oops text \A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp}\ lemma - "~ distinct ws ==> EX xs ys zs y. ws = xs @ [y] @ ys @ [y]" + "\ distinct ws \ \xs ys zs y. ws = xs @ [y] @ ys @ [y]" quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] oops text \A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last}\ lemma - "x : set xs ==> EX ys zs. xs = ys @ x # zs & x ~: set zs & x ~: set ys" + "x \ set xs \ \ys zs. xs = ys @ x # zs \ x \ set zs \ x \ set ys" quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] oops text \A false conjecture derived from @{thm map_eq_Cons_conv} with a typo\ lemma - "(map f xs = y # ys) = (EX z zs. xs = z' # zs & f z = y & map f zs = ys)" + "(map f xs = y # ys) = (\z zs. xs = z' # zs & f z = y \ map f zs = ys)" quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample] oops -lemma "a # xs = ys @ [a] ==> EX zs. xs = zs @ [a] & ys = a#zs" +lemma "a # xs = ys @ [a] \ \zs. xs = zs @ [a] \ ys = a#zs" quickcheck[narrowing, expect = counterexample] quickcheck[exhaustive, random, expect = no_counterexample] oops @@ -216,7 +216,7 @@ by (induct t) auto lemma is_ord_mkt: - "is_ord (MKT n l r h) = ((ALL n': set (set_of' l). n' < n) & (ALL n': set (set_of' r). n < n') & is_ord l & is_ord r)" + "is_ord (MKT n l r h) = ((\n' \ set (set_of' l). n' < n) \ (\n' \ set (set_of' r). n < n') \ is_ord l \ is_ord r)" by (simp add: set_of') declare is_ord.simps(1)[code] is_ord_mkt[code] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Cardholder_Registration.thy --- a/src/HOL/SET_Protocol/Cardholder_Registration.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy Thu Feb 15 13:04:36 2018 +0100 @@ -21,7 +21,7 @@ subsection\Predicate Formalizing the Encryption Association between Keys\ -primrec KeyCryptKey :: "[key, key, event list] => bool" +primrec KeyCryptKey :: "[key, key, event list] \ bool" where KeyCryptKey_Nil: "KeyCryptKey DK K [] = False" @@ -35,18 +35,18 @@ "KeyCryptKey DK K (ev # evs) = (KeyCryptKey DK K evs | (case ev of - Says A B Z => - ((\N X Y. A \ Spy & - DK \ symKeys & + Says A B Z \ + ((\N X Y. A \ Spy \ + DK \ symKeys \ Z = \Crypt DK \Agent A, Nonce N, Key K, X\, Y\) | (\C. DK = priEK C)) - | Gets A' X => False - | Notes A' X => False))" + | Gets A' X \ False + | Notes A' X \ False))" subsection\Predicate formalizing the association between keys and nonces\ -primrec KeyCryptNonce :: "[key, key, event list] => bool" +primrec KeyCryptNonce :: "[key, key, event list] \ bool" where KeyCryptNonce_Nil: "KeyCryptNonce EK K [] = False" @@ -63,25 +63,25 @@ "KeyCryptNonce DK N (ev # evs) = (KeyCryptNonce DK N evs | (case ev of - Says A B Z => - A \ Spy & - ((\X Y. DK \ symKeys & + Says A B Z \ + A \ Spy \ + ((\X Y. DK \ symKeys \ Z = (EXHcrypt DK X \Agent A, Nonce N\ Y)) | - (\X Y. DK \ symKeys & + (\X Y. DK \ symKeys \ Z = \Crypt DK \Agent A, Nonce N, X\, Y\) | (\K i X Y. - K \ symKeys & - Z = Crypt K \sign (priSK (CA i)) \Agent B, Nonce N, X\, Y\ & + K \ symKeys \ + Z = Crypt K \sign (priSK (CA i)) \Agent B, Nonce N, X\, Y\ \ (DK=K | KeyCryptKey DK K evs)) | (\K C NC3 Y. - K \ symKeys & + K \ symKeys \ Z = Crypt K \sign (priSK C) \Agent B, Nonce NC3, Agent C, Nonce N\, - Y\ & + Y\ \ (DK=K | KeyCryptKey DK K evs)) | (\C. DK = priEK C)) - | Gets A' X => False - | Notes A' X => False))" + | Gets A' X \ False + | Notes A' X \ False))" subsection\Formal protocol definition\ @@ -329,7 +329,7 @@ "[| Gets A \ X, cert C EKi onlyEnc (priSK RCA), cert C SKi onlySig (priSK RCA)\ \ set evs; evs \ set_cr |] - ==> EKi = pubEK C & SKi = pubSK C" + ==> EKi = pubEK C \ SKi = pubSK C" by (blast dest: certificate_valid_pubEK certificate_valid_pubSK) text\Nobody can have used non-existent keys!\ @@ -350,13 +350,13 @@ lemma gen_new_keys_not_used: "[|Key K \ used evs; K \ symKeys; evs \ set_cr |] - ==> Key K \ used evs --> K \ symKeys --> - K \ keysFor (parts (Key`KK Un knows Spy evs))" + ==> Key K \ used evs \ K \ symKeys \ + K \ keysFor (parts (Key`KK \ knows Spy evs))" by (auto simp add: new_keys_not_used) lemma gen_new_keys_not_analzd: "[|Key K \ used evs; K \ symKeys; evs \ set_cr |] - ==> K \ keysFor (analz (Key`KK Un knows Spy evs))" + ==> K \ keysFor (analz (Key`KK \ knows Spy evs))" by (blast intro: keysFor_mono [THEN [2] rev_subsetD] dest: gen_new_keys_not_used) @@ -459,14 +459,14 @@ for other keys aren't needed.\ lemma parts_image_priEK: - "[|Key (priEK C) \ parts (Key`KK Un (knows Spy evs)); + "[|Key (priEK C) \ parts (Key`KK \ (knows Spy evs)); evs \ set_cr|] ==> priEK C \ KK | C \ bad" by auto text\trivial proof because (priEK C) never appears even in (parts evs)\ lemma analz_image_priEK: "evs \ set_cr ==> - (Key (priEK C) \ analz (Key`KK Un (knows Spy evs))) = + (Key (priEK C) \ analz (Key`KK \ (knows Spy evs))) = (priEK C \ KK | C \ bad)" by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) @@ -478,7 +478,7 @@ text\A fresh DK cannot be associated with any other (with respect to a given trace).\ lemma DK_fresh_not_KeyCryptKey: - "[| Key DK \ used evs; evs \ set_cr |] ==> ~ KeyCryptKey DK K evs" + "[| Key DK \ used evs; evs \ set_cr |] ==> \ KeyCryptKey DK K evs" apply (erule rev_mp) apply (erule set_cr.induct) apply (simp_all (no_asm_simp)) @@ -489,7 +489,7 @@ DK isn't a private encryption key may be an artifact of the particular definition of KeyCryptKey.\ lemma K_fresh_not_KeyCryptKey: - "[|\C. DK \ priEK C; Key K \ used evs|] ==> ~ KeyCryptKey DK K evs" + "[|\C. DK \ priEK C; Key K \ used evs|] ==> \ KeyCryptKey DK K evs" apply (induct evs) apply (auto simp add: parts_insert2 split: event.split) done @@ -499,17 +499,17 @@ be known to the Spy, by @{term Spy_see_private_Key}\ lemma cardSK_neq_priEK: "[|Key cardSK \ analz (knows Spy evs); - Key cardSK : parts (knows Spy evs); + Key cardSK \ parts (knows Spy evs); evs \ set_cr|] ==> cardSK \ priEK C" by blast lemma not_KeyCryptKey_cardSK [rule_format (no_asm)]: "[|cardSK \ symKeys; \C. cardSK \ priEK C; evs \ set_cr|] ==> - Key cardSK \ analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs" + Key cardSK \ analz (knows Spy evs) \ \ KeyCryptKey cardSK K evs" by (erule set_cr.induct, analz_mono_contra, auto) text\Lemma for message 5: pubSK C is never used to encrypt Keys.\ -lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs" +lemma pubSK_not_KeyCryptKey [simp]: "\ KeyCryptKey (pubSK C) K evs" apply (induct_tac "evs") apply (auto simp add: parts_insert2 split: event.split) done @@ -523,14 +523,14 @@ \ set evs; cardSK \ symKeys; evs \ set_cr|] ==> Key cardSK \ analz (knows Spy evs) | - (\K. ~ KeyCryptKey cardSK K evs)" + (\K. \ KeyCryptKey cardSK K evs)" by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK) text\As usual: we express the property as a logical equivalence\ lemma Key_analz_image_Key_lemma: - "P --> (Key K \ analz (Key`KK Un H)) --> (K \ KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K \ KK | Key K \ analz H) ==> - P --> (Key K \ analz (Key`KK Un H)) = (K \ KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K \ KK | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) method_setup valid_certificate_tac = \ @@ -545,8 +545,8 @@ the quantifier and allows the simprule's condition to itself be simplified.\ lemma symKey_compromise [rule_format (no_asm)]: "evs \ set_cr ==> - (\SK KK. SK \ symKeys \ (\K \ KK. ~ KeyCryptKey K SK evs) --> - (Key SK \ analz (Key`KK Un (knows Spy evs))) = + (\SK KK. SK \ symKeys \ (\K \ KK. \ KeyCryptKey K SK evs) \ + (Key SK \ analz (Key`KK \ (knows Spy evs))) = (SK \ KK | Key SK \ analz (knows Spy evs)))" apply (erule set_cr.induct) apply (rule_tac [!] allI) + @@ -572,9 +572,9 @@ wrong!!\ lemma symKey_secrecy [rule_format]: "[|CA i \ bad; K \ symKeys; evs \ set_cr|] - ==> \X c. Says (Cardholder c) (CA i) X \ set evs --> - Key K \ parts{X} --> - Cardholder c \ bad --> + ==> \X c. Says (Cardholder c) (CA i) X \ set evs \ + Key K \ parts{X} \ + Cardholder c \ bad \ Key K \ analz (knows Spy evs)" apply (erule set_cr.induct) apply (frule_tac [8] Gets_certificate_valid) \ \for message 5\ @@ -629,7 +629,7 @@ lemma Hash_imp_parts [rule_format]: "evs \ set_cr - ==> Hash\X, Nonce N\ \ parts (knows Spy evs) --> + ==> Hash\X, Nonce N\ \ parts (knows Spy evs) \ Nonce N \ parts (knows Spy evs)" apply (erule set_cr.induct, force) apply (simp_all (no_asm_simp)) @@ -638,8 +638,8 @@ lemma Hash_imp_parts2 [rule_format]: "evs \ set_cr - ==> Hash\X, Nonce M, Y, Nonce N\ \ parts (knows Spy evs) --> - Nonce M \ parts (knows Spy evs) & Nonce N \ parts (knows Spy evs)" + ==> Hash\X, Nonce M, Y, Nonce N\ \ parts (knows Spy evs) \ + Nonce M \ parts (knows Spy evs) \ Nonce N \ parts (knows Spy evs)" apply (erule set_cr.induct, force) apply (simp_all (no_asm_simp)) apply (blast intro: parts_mono [THEN [2] rev_subsetD]) @@ -654,7 +654,7 @@ (with respect to a given trace).\ lemma DK_fresh_not_KeyCryptNonce: "[| DK \ symKeys; Key DK \ used evs; evs \ set_cr |] - ==> ~ KeyCryptNonce DK K evs" + ==> \ KeyCryptNonce DK K evs" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_cr.induct) @@ -667,7 +667,7 @@ text\A fresh N cannot be associated with any other (with respect to a given trace).\ lemma N_fresh_not_KeyCryptNonce: - "\C. DK \ priEK C ==> Nonce N \ used evs --> ~ KeyCryptNonce DK N evs" + "\C. DK \ priEK C ==> Nonce N \ used evs \ \ KeyCryptNonce DK N evs" apply (induct_tac "evs") apply (rename_tac [2] a evs') apply (case_tac [2] "a") @@ -676,7 +676,7 @@ lemma not_KeyCryptNonce_cardSK [rule_format (no_asm)]: "[|cardSK \ symKeys; \C. cardSK \ priEK C; evs \ set_cr|] ==> - Key cardSK \ analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs" + Key cardSK \ analz (knows Spy evs) \ \ KeyCryptNonce cardSK N evs" apply (erule set_cr.induct, analz_mono_contra, simp_all) apply (blast dest: not_KeyCryptKey_cardSK) \ \6\ done @@ -686,7 +686,7 @@ or else cardSK hasn't been used to encrypt K.\ text\Lemma for message 5: pubSK C is never used to encrypt Nonces.\ -lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs" +lemma pubSK_not_KeyCryptNonce [simp]: "\ KeyCryptNonce (pubSK C) N evs" apply (induct_tac "evs") apply (auto simp add: parts_insert2 split: event.split) done @@ -698,16 +698,16 @@ \ set evs; cardSK \ symKeys; evs \ set_cr|] ==> Key cardSK \ analz (knows Spy evs) | - ((\K. ~ KeyCryptKey cardSK K evs) & - (\N. ~ KeyCryptNonce cardSK N evs))" + ((\K. \ KeyCryptKey cardSK K evs) \ + (\N. \ KeyCryptNonce cardSK N evs))" by (blast dest: not_KeyCryptKey_cardSK not_KeyCryptNonce_cardSK intro: cardSK_neq_priEK) text\As usual: we express the property as a logical equivalence\ lemma Nonce_analz_image_Key_lemma: - "P --> (Nonce N \ analz (Key`KK Un H)) --> (Nonce N \ analz H) - ==> P --> (Nonce N \ analz (Key`KK Un H)) = (Nonce N \ analz H)" + "P \ (Nonce N \ analz (Key`KK \ H)) \ (Nonce N \ analz H) + ==> P \ (Nonce N \ analz (Key`KK \ H)) = (Nonce N \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) @@ -715,8 +715,8 @@ the quantifier and allows the simprule's condition to itself be simplified.\ lemma Nonce_compromise [rule_format (no_asm)]: "evs \ set_cr ==> - (\N KK. (\K \ KK. ~ KeyCryptNonce K N evs) --> - (Nonce N \ analz (Key`KK Un (knows Spy evs))) = + (\N KK. (\K \ KK. \ KeyCryptNonce K N evs) \ + (Nonce N \ analz (Key`KK \ (knows Spy evs))) = (Nonce N \ analz (knows Spy evs)))" apply (erule set_cr.induct) apply (rule_tac [!] allI)+ @@ -766,9 +766,9 @@ "[|U = Crypt KC3 \Agent C, Nonce N, Key KC2, X\; U \ parts (knows Spy evs); evs \ set_cr|] - ==> Nonce N \ analz (knows Spy evs) --> - (\k i W. Says (Cardholder k) (CA i) \U,W\ \ set evs & - Cardholder k \ bad & CA i \ bad)" + ==> Nonce N \ analz (knows Spy evs) \ + (\k i W. Says (Cardholder k) (CA i) \U,W\ \ set evs \ + Cardholder k \ bad \ CA i \ bad)" apply (erule_tac P = "U \ H" for H in rev_mp) apply (erule set_cr.induct) apply (valid_certificate_tac [8]) \ \for message 5\ @@ -791,9 +791,9 @@ text\Inductive version\ lemma CardSecret_secrecy_lemma [rule_format]: "[|CA i \ bad; evs \ set_cr|] - ==> Key K \ analz (knows Spy evs) --> + ==> Key K \ analz (knows Spy evs) \ Crypt (pubEK (CA i)) \Key K, Pan p, Nonce CardSecret\ - \ parts (knows Spy evs) --> + \ parts (knows Spy evs) \ Nonce CardSecret \ analz (knows Spy evs)" apply (erule set_cr.induct, analz_mono_contra) apply (valid_certificate_tac [8]) \ \for message 5\ @@ -854,12 +854,12 @@ text\Inductive version\ lemma NonceCCA_secrecy_lemma [rule_format]: "[|CA i \ bad; evs \ set_cr|] - ==> Key K \ analz (knows Spy evs) --> + ==> Key K \ analz (knows Spy evs) \ Crypt K \sign (priSK (CA i)) \Agent C, Nonce N, Agent(CA i), Nonce NonceCCA\, X, Y\ - \ parts (knows Spy evs) --> + \ parts (knows Spy evs) \ Nonce NonceCCA \ analz (knows Spy evs)" apply (erule set_cr.induct, analz_mono_contra) apply (valid_certificate_tac [8]) \ \for message 5\ @@ -921,14 +921,14 @@ by auto lemma analz_image_pan_lemma: - "(Pan P \ analz (Key`nE Un H)) --> (Pan P \ analz H) ==> - (Pan P \ analz (Key`nE Un H)) = (Pan P \ analz H)" + "(Pan P \ analz (Key`nE \ H)) \ (Pan P \ analz H) ==> + (Pan P \ analz (Key`nE \ H)) = (Pan P \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) lemma analz_image_pan [rule_format]: "evs \ set_cr ==> - \KK. KK <= - invKey ` pubEK ` range CA --> - (Pan P \ analz (Key`KK Un (knows Spy evs))) = + \KK. KK \ - invKey ` pubEK ` range CA \ + (Pan P \ analz (Key`KK \ (knows Spy evs))) = (Pan P \ analz (knows Spy evs))" apply (erule set_cr.induct) apply (rule_tac [!] allI impI)+ @@ -957,11 +957,11 @@ this theorem with @{term analz_image_pan}, requiring a single induction but a much more difficult proof.\ lemma pan_confidentiality: - "[| Pan (pan C) \ analz(knows Spy evs); C \Spy; evs :set_cr|] + "[| Pan (pan C) \ analz(knows Spy evs); C \Spy; evs \ set_cr|] ==> \i X K HN. Says C (CA i) \X, Crypt (pubEK (CA i)) \Key K, Pan (pan C), HN\ \ \ set evs - & (CA i) \ bad" + \ (CA i) \ bad" apply (erule rev_mp) apply (erule set_cr.induct) apply (valid_certificate_tac [8]) \ \for message 5\ @@ -1006,7 +1006,7 @@ certC (pan C') cardSK X' onlySig (priSK (CA i)), cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\) \ set evs; - evs \ set_cr |] ==> C=C' & NC3=NC3' & X=X' & KC2=KC2' & Y=Y'" + evs \ set_cr |] ==> C=C' \ NC3=NC3' \ X=X' \ KC2=KC2' \ Y=Y'" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_cr.induct) @@ -1022,7 +1022,7 @@ \ set evs; Says C B' \Crypt KC1 X', Crypt EK' \Key KC1, Y'\\ \ set evs; - C \ bad; evs \ set_cr|] ==> B'=B & Y'=Y" + C \ bad; evs \ set_cr|] ==> B'=B \ Y'=Y" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_cr.induct, auto) @@ -1032,7 +1032,7 @@ lemma unique_KC2: "[|Says C B \Crypt K \Agent C, nn, Key KC2, X\, Y\ \ set evs; Says C B' \Crypt K' \Agent C, nn', Key KC2, X'\, Y'\ \ set evs; - C \ bad; evs \ set_cr|] ==> B'=B & X'=X" + C \ bad; evs \ set_cr|] ==> B'=B \ X'=X" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_cr.induct, auto) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Event_SET.thy --- a/src/HOL/SET_Protocol/Event_SET.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Event_SET.thy Thu Feb 15 13:04:36 2018 +0100 @@ -34,11 +34,11 @@ subsection\Agents' Knowledge\ consts (*Initial states of agents -- parameter of the construction*) - initState :: "agent => msg set" + initState :: "agent \ msg set" (* Message reception does not extend spy's knowledge because of reception invariant enforced by Reception rule in protocol definition*) -primrec knows :: "[agent, event list] => msg set" +primrec knows :: "[agent, event list] \ msg set" where knows_Nil: "knows A [] = initState A" @@ -46,23 +46,23 @@ "knows A (ev # evs) = (if A = Spy then (case ev of - Says A' B X => insert X (knows Spy evs) - | Gets A' X => knows Spy evs - | Notes A' X => + Says A' B X \ insert X (knows Spy evs) + | Gets A' X \ knows Spy evs + | Notes A' X \ if A' \ bad then insert X (knows Spy evs) else knows Spy evs) else (case ev of - Says A' B X => + Says A' B X \ if A'=A then insert X (knows A evs) else knows A evs - | Gets A' X => + | Gets A' X \ if A'=A then insert X (knows A evs) else knows A evs - | Notes A' X => + | Notes A' X \ if A'=A then insert X (knows A evs) else knows A evs))" subsection\Used Messages\ -primrec used :: "event list => msg set" +primrec used :: "event list \ msg set" where (*Set of items that might be visible to somebody: complement of the set of fresh items. @@ -70,9 +70,9 @@ used_Nil: "used [] = (UN B. parts (initState B))" | used_Cons: "used (ev # evs) = (case ev of - Says A B X => parts {X} Un (used evs) - | Gets A X => used evs - | Notes A X => parts {X} Un (used evs))" + Says A B X \ parts {X} \ (used evs) + | Gets A X \ used evs + | Notes A X \ parts {X} \ (used evs))" @@ -80,7 +80,7 @@ be re-loaded. Addsimps [knows_Cons, used_Nil, *) (** Simplifying parts (insert X (knows Spy evs)) - = parts {X} Un parts (knows Spy evs) -- since general case loops*) + = parts {X} \ parts (knows Spy evs) -- since general case loops*) lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs @@ -92,33 +92,33 @@ on whether @{term "A=Spy"} and whether @{term "A\bad"}\ lemma knows_Spy_Notes [simp]: "knows Spy (Notes A X # evs) = - (if A:bad then insert X (knows Spy evs) else knows Spy evs)" + (if A\bad then insert X (knows Spy evs) else knows Spy evs)" apply auto done lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" by auto -lemma initState_subset_knows: "initState A <= knows A evs" +lemma initState_subset_knows: "initState A \ knows A evs" apply (induct_tac "evs") apply (auto split: event.split) done lemma knows_Spy_subset_knows_Spy_Says: - "knows Spy evs <= knows Spy (Says A B X # evs)" + "knows Spy evs \ knows Spy (Says A B X # evs)" by auto lemma knows_Spy_subset_knows_Spy_Notes: - "knows Spy evs <= knows Spy (Notes A X # evs)" + "knows Spy evs \ knows Spy (Notes A X # evs)" by auto lemma knows_Spy_subset_knows_Spy_Gets: - "knows Spy evs <= knows Spy (Gets A X # evs)" + "knows Spy evs \ knows Spy (Gets A X # evs)" by auto (*Spy sees what is sent on the traffic*) lemma Says_imp_knows_Spy [rule_format]: - "Says A B X \ set evs --> X \ knows Spy evs" + "Says A B X \ set evs \ X \ knows Spy evs" apply (induct_tac "evs") apply (auto split: event.split) done @@ -132,24 +132,24 @@ subsection\The Function @{term used}\ -lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) <= used evs" +lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \ used evs" apply (induct_tac "evs") apply (auto simp add: parts_insert_knows_A split: event.split) done lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] -lemma initState_subset_used: "parts (initState B) <= used evs" +lemma initState_subset_used: "parts (initState B) \ used evs" apply (induct_tac "evs") apply (auto split: event.split) done lemmas initState_into_used = initState_subset_used [THEN subsetD] -lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} Un used evs" +lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \ used evs" by auto -lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} Un used evs" +lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \ used evs" by auto lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" @@ -157,7 +157,7 @@ lemma Notes_imp_parts_subset_used [rule_format]: - "Notes A X \ set evs --> parts {X} <= used evs" + "Notes A X \ set evs \ parts {X} \ used evs" apply (induct_tac "evs") apply (rename_tac [2] a evs') apply (induct_tac [2] "a", auto) @@ -168,7 +168,7 @@ used_Nil [simp del] used_Cons [simp del] -text\For proving theorems of the form @{term "X \ analz (knows Spy evs) --> P"} +text\For proving theorems of the form @{term "X \ analz (knows Spy evs) \ P"} New events added by induction to "evs" are discarded. Provided this information isn't needed, the proof will be much shorter, since it will omit complicated reasoning about @{term analz}.\ @@ -190,6 +190,6 @@ method_setup analz_mono_contra = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\ - "for proving theorems of the form X \ analz (knows Spy evs) --> P" + "for proving theorems of the form X \ analz (knows Spy evs) \ P" end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Merchant_Registration.thy --- a/src/HOL/SET_Protocol/Merchant_Registration.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Merchant_Registration.thy Thu Feb 15 13:04:36 2018 +0100 @@ -182,14 +182,14 @@ "[| Gets A \ X, cert (CA i) EKi onlyEnc (priSK RCA), cert (CA i) SKi onlySig (priSK RCA)\ \ set evs; evs \ set_mr |] - ==> EKi = pubEK (CA i) & SKi = pubSK (CA i)" + ==> EKi = pubEK (CA i) \ SKi = pubSK (CA i)" by (blast dest: certificate_valid_pubEK certificate_valid_pubSK) text\Nobody can have used non-existent keys!\ lemma new_keys_not_used [rule_format,simp]: "evs \ set_mr - ==> Key K \ used evs --> K \ symKeys --> + ==> Key K \ used evs \ K \ symKeys \ K \ keysFor (parts (knows Spy evs))" apply (erule set_mr.induct, simp_all) apply (force dest!: usedI keysFor_parts_insert) \ \Fake\ @@ -203,13 +203,13 @@ lemma gen_new_keys_not_used [rule_format]: "evs \ set_mr - ==> Key K \ used evs --> K \ symKeys --> - K \ keysFor (parts (Key`KK Un knows Spy evs))" + ==> Key K \ used evs \ K \ symKeys \ + K \ keysFor (parts (Key`KK \ knows Spy evs))" by auto lemma gen_new_keys_not_analzd: "[|Key K \ used evs; K \ symKeys; evs \ set_mr |] - ==> K \ keysFor (analz (Key`KK Un knows Spy evs))" + ==> K \ keysFor (analz (Key`KK \ knows Spy evs))" by (blast intro: keysFor_mono [THEN [2] rev_subsetD] dest: gen_new_keys_not_used) @@ -236,14 +236,14 @@ for other keys aren't needed.\ lemma parts_image_priEK: - "[|Key (priEK (CA i)) \ parts (Key`KK Un (knows Spy evs)); + "[|Key (priEK (CA i)) \ parts (Key`KK \ (knows Spy evs)); evs \ set_mr|] ==> priEK (CA i) \ KK | CA i \ bad" by auto text\trivial proof because (priEK (CA i)) never appears even in (parts evs)\ lemma analz_image_priEK: "evs \ set_mr ==> - (Key (priEK (CA i)) \ analz (Key`KK Un (knows Spy evs))) = + (Key (priEK (CA i)) \ analz (Key`KK \ (knows Spy evs))) = (priEK (CA i) \ KK | CA i \ bad)" by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) @@ -266,22 +266,22 @@ Y\ \ set evs; evs \ set_mr|] ==> (Key merSK \ analz (knows Spy evs) | merSK \ range(\C. priEK C)) - & (Key merEK \ analz (knows Spy evs) | merEK \ range(\C. priEK C))" + \ (Key merEK \ analz (knows Spy evs) | merEK \ range(\C. priEK C))" apply (unfold sign_def) apply (blast dest: merK_neq_priEK) done lemma Key_analz_image_Key_lemma: - "P --> (Key K \ analz (Key`KK Un H)) --> (K\KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K\KK | Key K \ analz H) ==> - P --> (Key K \ analz (Key`KK Un H)) = (K\KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K\KK | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) lemma symKey_compromise: "evs \ set_mr ==> - (\SK KK. SK \ symKeys \ (\K \ KK. K \ range(\C. priEK C)) --> - (Key SK \ analz (Key`KK Un (knows Spy evs))) = + (\SK KK. SK \ symKeys \ (\K \ KK. K \ range(\C. priEK C)) \ + (Key SK \ analz (Key`KK \ (knows Spy evs))) = (SK \ KK | Key SK \ analz (knows Spy evs)))" apply (erule set_mr.induct) apply (safe del: impI intro!: Key_analz_image_Key_lemma [THEN impI]) @@ -299,9 +299,9 @@ lemma symKey_secrecy [rule_format]: "[|CA i \ bad; K \ symKeys; evs \ set_mr|] - ==> \X m. Says (Merchant m) (CA i) X \ set evs --> - Key K \ parts{X} --> - Merchant m \ bad --> + ==> \X m. Says (Merchant m) (CA i) X \ set evs \ + Key K \ parts{X} \ + Merchant m \ bad \ Key K \ analz (knows Spy evs)" apply (erule set_mr.induct) apply (drule_tac [7] msg4_priEK_disj) @@ -326,7 +326,7 @@ cert (CA i) (pubSK (CA i)) onlySig (priSK RCA)\ \ set evs; evs \ set_mr |] ==> Notes (CA i) (Key merSK) \ set evs - & Notes (CA i) (Key merEK) \ set evs" + \ Notes (CA i) (Key merEK) \ set evs" apply (erule rev_mp) apply (erule set_mr.induct) apply (simp_all (no_asm_simp)) @@ -343,7 +343,7 @@ cert M' merSK onlySig (priSK (CA i)), cert M' merEK' onlyEnc (priSK (CA i)), cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)\ \ set evs; - evs \ set_mr |] ==> M=M' & NM2=NM2' & merEK=merEK'" + evs \ set_mr |] ==> M=M' \ NM2=NM2' \ merEK=merEK'" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_mr.induct) @@ -363,7 +363,7 @@ cert M' merEK onlyEnc (priSK (CA i)), cert (CA i) (pubSK(CA i)) onlySig (priSK RCA)\ \ set evs; evs \ set_mr |] - ==> M=M' & NM2=NM2' & merSK=merSK'" + ==> M=M' \ NM2=NM2' \ merSK=merSK'" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_mr.induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Message_SET.thy --- a/src/HOL/SET_Protocol/Message_SET.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Message_SET.thy Thu Feb 15 13:04:36 2018 +0100 @@ -27,7 +27,7 @@ by blast text\Effective with the assumption @{term "KK \ - (range(invKey o pubK))"}\ -lemma disjoint_image_iff: "(A <= - (f`I)) = (\i\I. f i \ A)" +lemma disjoint_image_iff: "(A \ - (f`I)) = (\i\I. f i \ A)" by blast @@ -36,11 +36,11 @@ consts all_symmetric :: bool \ \true if all keys are symmetric\ - invKey :: "key=>key" \ \inverse of a symmetric key\ + invKey :: "key\key" \ \inverse of a symmetric key\ specification (invKey) invKey [simp]: "invKey (invKey K) = K" - invKey_symmetric: "all_symmetric --> invKey = id" + invKey_symmetric: "all_symmetric \ invKey = id" by (rule exI [of _ id], auto) @@ -69,13 +69,13 @@ (*Concrete syntax: messages appear as \\A,B,NA\\, etc...*) syntax - "_MTuple" :: "['a, args] => 'a * 'b" ("(2\_,/ _\)") + "_MTuple" :: "['a, args] \ 'a * 'b" ("(2\_,/ _\)") translations "\x, y, z\" == "\x, \y, z\\" "\x, y\" == "CONST MPair x y" -definition nat_of_agent :: "agent => nat" where +definition nat_of_agent :: "agent \ nat" where "nat_of_agent == case_agent (curry prod_encode 0) (curry prod_encode 1) (curry prod_encode 2) @@ -90,13 +90,13 @@ definition (*Keys useful to decrypt elements of a message set*) - keysFor :: "msg set => key set" + keysFor :: "msg set \ key set" where "keysFor H = invKey ` {K. \X. Crypt K X \ H}" subsubsection\Inductive definition of all "parts" of a message.\ inductive_set - parts :: "msg set => msg set" + parts :: "msg set \ msg set" for H :: "msg set" where Inj [intro]: "X \ H ==> X \ parts H" @@ -106,7 +106,7 @@ (*Monotonicity*) -lemma parts_mono: "G<=H ==> parts(G) <= parts(H)" +lemma parts_mono: "G\H ==> parts(G) \ parts(H)" apply auto apply (erule parts.induct) apply (auto dest: Fst Snd Body) @@ -363,7 +363,7 @@ (*In any message, there is an upper bound N on its greatest nonce.*) -lemma msg_Nonce_supply: "\N. \n. N\n --> Nonce n \ parts {msg}" +lemma msg_Nonce_supply: "\N. \n. N\n \ Nonce n \ parts {msg}" apply (induct_tac "msg") apply (simp_all (no_asm_simp) add: exI parts_insert2) (*MPair case: blast_tac works out the necessary sum itself!*) @@ -375,7 +375,7 @@ done (* Ditto, for numbers.*) -lemma msg_Number_supply: "\N. \n. N<=n --> Number n \ parts {msg}" +lemma msg_Number_supply: "\N. \n. N\n \ Number n \ parts {msg}" apply (induct_tac "msg") apply (simp_all (no_asm_simp) add: exI parts_insert2) prefer 2 apply (blast elim!: add_leE) @@ -397,11 +397,11 @@ | 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" + "[|Crypt K X \ analz H; Key(invKey K) \ analz H|] ==> X \ analz H" (*Monotonicity; Lemma 1 of Lowe's paper*) -lemma analz_mono: "G<=H ==> analz(G) <= analz(H)" +lemma analz_mono: "G\H ==> analz(G) \ analz(H)" apply auto apply (erule analz.induct) apply (auto dest: Fst Snd) @@ -584,7 +584,7 @@ by (erule analz_trans, blast) (*Cut can be proved easily by induction on - "Y: analz (insert X H) ==> X: analz H --> Y: analz H" + "Y: analz (insert X H) ==> X: analz H \ Y: analz H" *) (*This rewrite rule helps in the simplification of messages that involve @@ -639,7 +639,7 @@ Numbers can be guessed, but Nonces cannot be.\ inductive_set - synth :: "msg set => msg set" + synth :: "msg set \ msg set" for H :: "msg set" where Inj [intro]: "X \ H ==> X \ synth H" @@ -650,7 +650,7 @@ | Crypt [intro]: "[|X \ synth H; Key(K) \ H|] ==> Crypt K X \ synth H" (*Monotonicity*) -lemma synth_mono: "G<=H ==> synth(G) <= synth(H)" +lemma synth_mono: "G\H ==> synth(G) \ synth(H)" apply auto apply (erule synth.induct) apply (auto dest: Fst Snd Body) @@ -760,7 +760,7 @@ done lemma Fake_parts_insert_in_Un: - "[|Z \ parts (insert X H); X: synth (analz H)|] + "[|Z \ parts (insert X H); X \ synth (analz H)|] ==> Z \ synth (analz H) \ parts H" by (blast dest: Fake_parts_insert [THEN subsetD, dest]) @@ -776,7 +776,7 @@ done lemma analz_conj_parts [simp]: - "(X \ analz H & X \ parts H) = (X \ analz H)" + "(X \ analz H \ X \ parts H) = (X \ analz H)" by (blast intro: analz_subset_parts [THEN subsetD]) lemma analz_disj_parts [simp]: @@ -787,7 +787,7 @@ redundant cases*) lemma MPair_synth_analz [iff]: "(\X,Y\ \ synth (analz H)) = - (X \ synth (analz H) & Y \ synth (analz H))" + (X \ synth (analz H) \ Y \ synth (analz H))" by blast lemma Crypt_synth_analz: @@ -889,7 +889,7 @@ lemma Hash_notin_image_Key [simp] :"Hash X \ Key ` A" by auto -lemma synth_analz_mono: "G<=H ==> synth (analz(G)) <= synth (analz(H))" +lemma synth_analz_mono: "G\H ==> synth (analz(G)) \ synth (analz(H))" by (simp add: synth_mono analz_mono) lemma Fake_analz_eq [simp]: @@ -903,12 +903,12 @@ text\Two generalizations of \analz_insert_eq\\ lemma gen_analz_insert_eq [rule_format]: - "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G" + "X \ analz H ==> \G. H \ G \ analz (insert X G) = analz G" by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) lemma synth_analz_insert_eq [rule_format]: "X \ synth (analz H) - ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)" + \ \G. H \ G \ (Key K \ analz (insert X G)) = (Key K \ analz G)" apply (erule synth.induct) apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Public_SET.thy --- a/src/HOL/SET_Protocol/Public_SET.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Public_SET.thy Thu Feb 15 13:04:36 2018 +0100 @@ -20,7 +20,7 @@ text\The SET specs mention two signature keys for CAs - we only have one\ consts - publicKey :: "[bool, agent] => key" + publicKey :: "[bool, agent] \ key" \ \the boolean is TRUE if a signing key\ abbreviation "pubEK == publicKey False" @@ -35,7 +35,7 @@ specification (publicKey) injective_publicKey: - "publicKey b A = publicKey c A' ==> b=c & A=A'" + "publicKey b A = publicKey c A' \ b=c \ A=A'" (*<*) apply (rule exI [of _ "%b A. 2 * nat_of_agent A + (if b then 1 else 0)"]) apply (auto simp add: inj_on_def inj_nat_of_agent [THEN inj_eq] split: agent.split) @@ -72,8 +72,8 @@ (*<*) initState_CA: "initState (CA i) = - (if i=0 then Key ` ({priEK RCA, priSK RCA} Un - pubEK ` (range CA) Un pubSK ` (range CA)) + (if i=0 then Key ` ({priEK RCA, priSK RCA} \ + pubEK ` (range CA) \ pubSK ` (range CA)) else {Key (priEK (CA i)), Key (priSK (CA i)), Key (pubEK (CA i)), Key (pubSK (CA i)), Key (pubEK RCA), Key (pubSK RCA)})" @@ -97,16 +97,16 @@ Key (pubEK RCA), Key (pubSK RCA)}" (*>*) | initState_Spy: - "initState Spy = Key ` (invKey ` pubEK ` bad Un - invKey ` pubSK ` bad Un - range pubEK Un range pubSK)" + "initState Spy = Key ` (invKey ` pubEK ` bad \ + invKey ` pubSK ` bad \ + range pubEK \ range pubSK)" end text\Injective mapping from agents to PANs: an agent can have only one card\ -consts pan :: "agent => nat" +consts pan :: "agent \ nat" specification (pan) inj_pan: "inj pan" @@ -120,24 +120,24 @@ declare inj_pan [THEN inj_eq, iff] consts - XOR :: "nat*nat => nat" \ \no properties are assumed of exclusive-or\ + XOR :: "nat*nat \ nat" \ \no properties are assumed of exclusive-or\ subsection\Signature Primitives\ definition (* Signature = Message + signed Digest *) - sign :: "[key, msg]=>msg" + sign :: "[key, msg]\msg" where "sign K X = \X, Crypt K (Hash X) \" definition (* Signature Only = signed Digest Only *) - signOnly :: "[key, msg]=>msg" + signOnly :: "[key, msg]\msg" where "signOnly K X = Crypt K (Hash X)" definition (* Signature for Certificates = Message + signed Message *) - signCert :: "[key, msg]=>msg" + signCert :: "[key, msg]\msg" where "signCert K X = \X, Crypt K X \" definition @@ -148,7 +148,7 @@ Should prove if signK=priSK RCA and C=CA i, then Ka=pubEK i or pubSK i depending on T ?? *) - cert :: "[agent, key, msg, key] => msg" + cert :: "[agent, key, msg, key] \ msg" where "cert A Ka T signK = signCert signK \Agent A, Key Ka, T\" definition @@ -156,7 +156,7 @@ Contains a PAN, the certified key Ka, the PANSecret PS, a number specifying the target use for Ka, the signing key signK. *) - certC :: "[nat, key, nat, msg, key] => msg" + certC :: "[nat, key, nat, msg, key] \ msg" where "certC PAN Ka PS T signK = signCert signK \Hash \Nonce PS, Pan PAN\, Key Ka, T\" @@ -169,19 +169,19 @@ subsection\Encryption Primitives\ -definition EXcrypt :: "[key,key,msg,msg] => msg" where +definition EXcrypt :: "[key,key,msg,msg] \ msg" where \ \Extra Encryption\ (*K: the symmetric key EK: the public encryption key*) "EXcrypt K EK M m = \Crypt K \M, Hash m\, Crypt EK \Key K, m\\" -definition EXHcrypt :: "[key,key,msg,msg] => msg" where +definition EXHcrypt :: "[key,key,msg,msg] \ msg" where \ \Extra Encryption with Hashing\ (*K: the symmetric key EK: the public encryption key*) "EXHcrypt K EK M m = \Crypt K \M, Hash m\, Crypt EK \Key K, m, Hash M\\" -definition Enc :: "[key,key,key,msg] => msg" where +definition Enc :: "[key,key,key,msg] \ msg" where \ \Simple Encapsulation with SIGNATURE\ (*SK: the sender's signing key K: the symmetric key @@ -189,7 +189,7 @@ "Enc SK K EK M = \Crypt K (sign SK M), Crypt EK (Key K)\" -definition EncB :: "[key,key,key,msg,msg] => msg" where +definition EncB :: "[key,key,key,msg,msg] \ msg" where \ \Encapsulation with Baggage. Keys as above, and baggage b.\ "EncB SK K EK M b = \Enc SK K EK \M, Hash b\, b\" @@ -198,11 +198,11 @@ subsection\Basic Properties of pubEK, pubSK, priEK and priSK\ lemma publicKey_eq_iff [iff]: - "(publicKey b A = publicKey b' A') = (b=b' & A=A')" + "(publicKey b A = publicKey b' A') = (b=b' \ A=A')" by (blast dest: injective_publicKey) lemma privateKey_eq_iff [iff]: - "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' & A=A')" + "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' \ A=A')" by auto lemma not_symKeys_publicKey [iff]: "publicKey b A \ symKeys" @@ -211,28 +211,28 @@ lemma not_symKeys_privateKey [iff]: "invKey (publicKey b A) \ symKeys" by (simp add: symKeys_def) -lemma symKeys_invKey_eq [simp]: "K \ symKeys ==> invKey K = K" +lemma symKeys_invKey_eq [simp]: "K \ symKeys \ invKey K = K" by (simp add: symKeys_def) lemma symKeys_invKey_iff [simp]: "(invKey K \ symKeys) = (K \ symKeys)" by (unfold symKeys_def, auto) text\Can be slow (or even loop) as a simprule\ -lemma symKeys_neq_imp_neq: "(K \ symKeys) \ (K' \ symKeys) ==> K \ K'" +lemma symKeys_neq_imp_neq: "(K \ symKeys) \ (K' \ symKeys) \ K \ K'" by blast text\These alternatives to \symKeys_neq_imp_neq\ don't seem any better in practice.\ -lemma publicKey_neq_symKey: "K \ symKeys ==> publicKey b A \ K" +lemma publicKey_neq_symKey: "K \ symKeys \ publicKey b A \ K" by blast -lemma symKey_neq_publicKey: "K \ symKeys ==> K \ publicKey b A" +lemma symKey_neq_publicKey: "K \ symKeys \ K \ publicKey b A" by blast -lemma privateKey_neq_symKey: "K \ symKeys ==> invKey (publicKey b A) \ K" +lemma privateKey_neq_symKey: "K \ symKeys \ invKey (publicKey b A) \ K" by blast -lemma symKey_neq_privateKey: "K \ symKeys ==> K \ invKey (publicKey b A)" +lemma symKey_neq_privateKey: "K \ symKeys \ K \ invKey (publicKey b A)" by blast lemma analz_symKeys_Decrypt: @@ -248,11 +248,11 @@ text\holds because invKey is injective\ lemma publicKey_image_eq [iff]: - "(publicKey b A \ publicKey c ` AS) = (b=c & A\AS)" + "(publicKey b A \ publicKey c ` AS) = (b=c \ A\AS)" by auto lemma privateKey_image_eq [iff]: - "(invKey (publicKey b A) \ invKey ` publicKey c ` AS) = (b=c & A\AS)" + "(invKey (publicKey b A) \ invKey ` publicKey c ` AS) = (b=c \ A\AS)" by auto lemma privateKey_notin_image_publicKey [iff]: @@ -308,7 +308,7 @@ text\Spy sees private keys of bad agents! [and obviously public keys too]\ lemma knows_Spy_bad_privateKey [intro!]: - "A \ bad ==> Key (invKey (publicKey b A)) \ knows Spy evs" + "A \ bad \ Key (invKey (publicKey b A)) \ knows Spy evs" by (rule initState_subset_knows [THEN subsetD], simp) @@ -321,7 +321,7 @@ by (simp add: used_Nil) text\In any trace, there is an upper bound N on the greatest nonce in use.\ -lemma Nonce_supply_lemma: "\N. \n. N<=n --> Nonce n \ used evs" +lemma Nonce_supply_lemma: "\N. \n. N\n \ Nonce n \ used evs" apply (induct_tac "evs") apply (rule_tac x = 0 in exI) apply (simp_all add: used_Cons split: event.split, safe) @@ -331,7 +331,7 @@ lemma Nonce_supply1: "\N. Nonce N \ used evs" by (rule Nonce_supply_lemma [THEN exE], blast) -lemma Nonce_supply: "Nonce (@ N. Nonce N \ used evs) \ used evs" +lemma Nonce_supply: "Nonce (SOME N. Nonce N \ used evs) \ used evs" apply (rule Nonce_supply_lemma [THEN exE]) apply (rule someI, fast) done @@ -369,11 +369,11 @@ subsection\Specialized Rewriting for Theorems About @{term analz} and Image\ -lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H" +lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \ H" by blast lemma insert_Key_image: - "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C" + "insert (Key K) (Key`KK \ C) = Key ` (insert K KK) \ C" by blast text\Needed for \DK_fresh_not_KeyCryptKey\\ @@ -398,16 +398,16 @@ text\A set is expanded only if a relation is applied to it\ lemma def_abbrev_simp_relation: - "A = B ==> (A \ X) = (B \ X) & - (u = A) = (u = B) & + "A = B \ (A \ X) = (B \ X) \ + (u = A) = (u = B) \ (A = u) = (B = u)" by auto text\A set is expanded only if one of the given functions is applied to it\ lemma def_abbrev_simp_function: "A = B - ==> parts (insert A X) = parts (insert B X) & - analz (insert A X) = analz (insert B X) & + \ parts (insert A X) = parts (insert B X) \ + analz (insert A X) = analz (insert B X) \ keysFor (insert A X) = keysFor (insert B X)" by auto @@ -516,7 +516,7 @@ by auto lemma fresh_notin_analz_knows_Spy: - "Key K \ used evs ==> Key K \ analz (knows Spy evs)" + "Key K \ used evs \ Key K \ analz (knows Spy evs)" by (auto dest: analz_into_parts) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SET_Protocol/Purchase.thy --- a/src/HOL/SET_Protocol/Purchase.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SET_Protocol/Purchase.thy Thu Feb 15 13:04:36 2018 +0100 @@ -53,11 +53,11 @@ consts - CardSecret :: "nat => nat" + CardSecret :: "nat \ nat" \ \Maps Cardholders to CardSecrets. A CardSecret of 0 means no cerificate, must use unsigned format.\ - PANSecret :: "nat => nat" + PANSecret :: "nat \ nat" \ \Maps Cardholders to PANSecrets.\ inductive_set @@ -192,7 +192,7 @@ HOD = Hash\Number OrderDesc, Number PurchAmt\; OIData = \Number LID_M, Number XID, Nonce Chall_C, HOD, Nonce Chall_M\; - CardSecret k \ 0 --> + CardSecret k \ 0 \ P_I = \sign (priSK C) \HPIData, Hash OIData\, encPANData\; Gets M \P_I, OIData, HPIData\ \ set evsAReq; Says M C (sign (priSK M) \Number LID_M, Number XID, @@ -427,7 +427,7 @@ text\rewriting rule for priEK's\ lemma parts_image_priEK: - "[|Key (priEK C) \ parts (Key`KK Un (knows Spy evs)); + "[|Key (priEK C) \ parts (Key`KK \ (knows Spy evs)); evs \ set_pur|] ==> priEK C \ KK | C \ bad" by auto @@ -435,7 +435,7 @@ @{term "parts evs"}.\ lemma analz_image_priEK: "evs \ set_pur ==> - (Key (priEK C) \ analz (Key`KK Un (knows Spy evs))) = + (Key (priEK C) \ analz (Key`KK \ (knows Spy evs))) = (priEK C \ KK | C \ bad)" by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD]) @@ -492,7 +492,7 @@ text\Nobody can have used non-existent keys!\ lemma new_keys_not_used [rule_format,simp]: "evs \ set_pur - ==> Key K \ used evs --> K \ symKeys --> + ==> Key K \ used evs \ K \ symKeys \ K \ keysFor (parts (knows Spy evs))" apply (erule set_pur.induct) apply (valid_certificate_tac [8]) \ \PReqS\ @@ -522,13 +522,13 @@ lemma gen_new_keys_not_used: "[|Key K \ used evs; K \ symKeys; evs \ set_pur |] - ==> Key K \ used evs --> K \ symKeys --> - K \ keysFor (parts (Key`KK Un knows Spy evs))" + ==> Key K \ used evs \ K \ symKeys \ + K \ keysFor (parts (Key`KK \ knows Spy evs))" by auto lemma gen_new_keys_not_analzd: "[|Key K \ used evs; K \ symKeys; evs \ set_pur |] - ==> K \ keysFor (analz (Key`KK Un knows Spy evs))" + ==> K \ keysFor (analz (Key`KK \ knows Spy evs))" by (blast intro: keysFor_mono [THEN subsetD] dest: gen_new_keys_not_used) lemma analz_Key_image_insert_eq: @@ -541,9 +541,9 @@ subsection\Secrecy of Symmetric Keys\ lemma Key_analz_image_Key_lemma: - "P --> (Key K \ analz (Key`KK Un H)) --> (K\KK | Key K \ analz H) + "P \ (Key K \ analz (Key`KK \ H)) \ (K\KK | Key K \ analz H) ==> - P --> (Key K \ analz (Key`KK Un H)) = (K\KK | Key K \ analz H)" + P \ (Key K \ analz (Key`KK \ H)) = (K\KK | Key K \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) @@ -575,15 +575,15 @@ text\As usual: we express the property as a logical equivalence\ lemma Nonce_analz_image_Key_lemma: - "P --> (Nonce N \ analz (Key`KK Un H)) --> (Nonce N \ analz H) - ==> P --> (Nonce N \ analz (Key`KK Un H)) = (Nonce N \ analz H)" + "P \ (Nonce N \ analz (Key`KK \ H)) \ (Nonce N \ analz H) + ==> P \ (Nonce N \ analz (Key`KK \ H)) = (Nonce N \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) text\The \(no_asm)\ attribute is essential, since it retains the quantifier and allows the simprule's condition to itself be simplified.\ lemma Nonce_compromise [rule_format (no_asm)]: "evs \ set_pur ==> - (\N KK. (\K \ KK. K \ range(\C. priEK C)) --> + (\N KK. (\K \ KK. K \ range(\C. priEK C)) \ (Nonce N \ analz (Key`KK \ (knows Spy evs))) = (Nonce N \ analz (knows Spy evs)))" apply (erule set_pur.induct) @@ -639,16 +639,16 @@ subsection\Confidentiality of PAN\ lemma analz_image_pan_lemma: - "(Pan P \ analz (Key`nE Un H)) --> (Pan P \ analz H) ==> - (Pan P \ analz (Key`nE Un H)) = (Pan P \ analz H)" + "(Pan P \ analz (Key`nE \ H)) \ (Pan P \ analz H) ==> + (Pan P \ analz (Key`nE \ H)) = (Pan P \ analz H)" by (blast intro: analz_mono [THEN [2] rev_subsetD]) text\The \(no_asm)\ attribute is essential, since it retains the quantifier and allows the simprule's condition to itself be simplified.\ lemma analz_image_pan [rule_format (no_asm)]: "evs \ set_pur ==> - \KK. (\K \ KK. K \ range(\C. priEK C)) --> - (Pan P \ analz (Key`KK Un (knows Spy evs))) = + \KK. (\K \ KK. K \ range(\C. priEK C)) \ + (Pan P \ analz (Key`KK \ (knows Spy evs))) = (Pan P \ analz (knows Spy evs))" apply (erule set_pur.induct) apply (rule_tac [!] allI impI)+ @@ -680,7 +680,7 @@ CardSecret k = 0; evs \ set_pur|] ==> \P M KC1 K X Y. Says C M \EXHcrypt KC1 (pubEK P) X (Pan (pan C)), Y\ - \ set evs & + \ set evs \ P \ bad" apply (erule rev_mp) apply (erule set_pur.induct) @@ -705,7 +705,7 @@ ==> \P M KC2 PIDualSign_1 PIDualSign_2 other OIDualSign. Says C M \\PIDualSign_1, EXcrypt KC2 (pubEK P) PIDualSign_2 \Pan (pan C), other\\, - OIDualSign\ \ set evs & P \ bad" + OIDualSign\ \ set evs \ P \ bad" apply (erule rev_mp) apply (erule set_pur.induct) apply (frule_tac [9] AuthReq_msg_in_analz_spies) \ \AReq\ @@ -742,7 +742,7 @@ Crypt (priSK M) (Hash MsgPInitRes) \ parts (knows Spy evs); evs \ set_pur; M \ bad |] ==> \j trans. - P = PG j & + P = PG j \ Notes M \Number LID_M, Agent P, trans\ \ set evs" apply clarify apply (erule rev_mp) @@ -757,8 +757,8 @@ cert P EKj onlyEnc (priSK RCA)\) \ set evs; evs \ set_pur; M \ bad|] ==> \j trans. - P = PG j & - Notes M \Number LID_M, Agent P, trans\ \ set evs & + P = PG j \ + Notes M \Number LID_M, Agent P, trans\ \ set evs \ EKj = pubEK P" by (rule refl [THEN goodM_gives_correct_PG, THEN exE], auto) @@ -769,8 +769,8 @@ Crypt (priSK M) (Hash MsgPInitRes) \ parts (knows Spy evs); evs \ set_pur; M \ bad|] ==> \j trans. - Notes M \Number LID_M, Agent P, trans\ \ set evs & - P = PG j & + Notes M \Number LID_M, Agent P, trans\ \ set evs \ + P = PG j \ EKj = pubEK P" apply clarify apply (erule rev_mp) @@ -788,8 +788,8 @@ cert P EKj onlyEnc (priSK RCA)\) \ set evs; M \ bad; evs \ set_pur|] ==> \j trans. - Notes M \Number LID_M, Agent P, trans\ \ set evs & - P = PG j & + Notes M \Number LID_M, Agent P, trans\ \ set evs \ + P = PG j \ EKj = pubEK (PG j)" apply (frule Says_certificate_valid) apply (auto simp add: sign_def) @@ -805,8 +805,8 @@ \ parts (knows Spy evs); evs \ set_pur; M \ bad|] ==> \j trans KM OIData HPIData. - Notes M \Number LID_M, Agent (PG j), trans\ \ set evs & - Gets M \P_I, OIData, HPIData\ \ set evs & + Notes M \Number LID_M, Agent (PG j), trans\ \ set evs \ + Gets M \P_I, OIData, HPIData\ \ set evs \ Says M (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I) \ set evs" apply clarify @@ -833,7 +833,7 @@ Gets (PG j) (EncB (priSK M) KM (pubEK (PG j)) \Number LID_M, Number XID, HOIData, HOD\ - P_I) \ set evs & + P_I) \ set evs \ Says (PG j) M (EncB (priSK (PG j)) KP (pubEK M) \Number LID_M, Number XID, Number PurchAmt\ @@ -857,7 +857,7 @@ PIHead = \Number LID_M, Trans_details\; evs \ set_pur; C = Cardholder k; M \ bad|] ==> \trans j. - Notes M \Number LID_M, Agent (PG j), trans \ \ set evs & + Notes M \Number LID_M, Agent (PG j), trans \ \ set evs \ EKj = pubEK (PG j)" apply clarify apply (erule rev_mp) @@ -874,7 +874,7 @@ Notes C \Number LID_M, Agent M, Agent C, Number OD, Number PA\ \ set evs; evs \ set_pur|] - ==> M = Merchant i & Trans = \Agent M, Agent C, Number OD, Number PA\" + ==> M = Merchant i \ Trans = \Agent M, Agent C, Number OD, Number PA\" apply (erule rev_mp) apply (erule rev_mp) apply (erule set_pur.induct, simp_all) @@ -935,11 +935,11 @@ evs \ set_pur; M \ bad|] ==> \j KP. Notes M \Number LID_M, Agent (PG j), Trans \ - \ set evs & + \ set evs \ Gets M (EncB (priSK (PG j)) KP (pubEK M) \Number LID_M, Number XID, Number PurchAmt\ authCode) - \ set evs & + \ set evs \ Says M C (sign (priSK M) MsgPRes) \ set evs" apply clarify apply (erule rev_mp) @@ -965,10 +965,10 @@ Number PurchAmt\ \ set evs; evs \ set_pur; M \ bad|] ==> \P KP trans. - Notes M \Number LID_M,Agent P, trans\ \ set evs & + Notes M \Number LID_M,Agent P, trans\ \ set evs \ Gets M (EncB (priSK P) KP (pubEK M) \Number LID_M, Number XID, Number PurchAmt\ - authCode) \ set evs & + authCode) \ set evs \ Says M C (sign (priSK M) MsgPRes) \ set evs" apply (rule C_verifies_PRes_lemma [THEN exE]) apply (auto simp add: sign_def) @@ -1034,7 +1034,7 @@ Notes M \Number LID_M, Agent P, extras\ \ set evs; M = Merchant i; C = Cardholder k; C \ bad; evs \ set_pur|] ==> \PIData PICrypt. - HPIData = Hash PIData & + HPIData = Hash PIData \ Says C M \\sign (priSK C) MsgDualSign, PICrypt\, OIData, Hash PIData\ \ set evs" apply clarify @@ -1061,9 +1061,9 @@ Crypt (priSK C) (Hash MsgDualSign) \ parts (knows Spy evs); evs \ set_pur; C \ bad; M \ bad|] ==> \OIData OrderDesc K j trans. - HOD = Hash\Number OrderDesc, Number PurchAmt\ & - HOIData = Hash OIData & - Notes M \Number LID_M, Agent (PG j), trans\ \ set evs & + HOD = Hash\Number OrderDesc, Number PurchAmt\ \ + HOIData = Hash OIData \ + Notes M \Number LID_M, Agent (PG j), trans\ \ set evs \ Says C M \\sign (priSK C) MsgDualSign, EXcrypt K (pubEK (PG j)) \PIHead, Hash OIData\ PANData\, @@ -1081,7 +1081,7 @@ PIHead = \Number LID_M, Number XID, W\; C = Cardholder k; evs \ set_pur; M \ bad|] ==> \ trans j. - Notes M \Number LID_M, Agent (PG j), trans\ \ set evs & + Notes M \Number LID_M, Agent (PG j), trans\ \ set evs \ EKj = pubEK (PG j)" apply clarify apply (erule rev_mp) @@ -1094,7 +1094,7 @@ sign (priSK M) \AuthReqData, Hash P_I\ \ parts (knows Spy evs); evs \ set_pur; M \ bad|] ==> \j trans KM. - Notes M \Number LID_M, Agent (PG j), trans \ \ set evs & + Notes M \Number LID_M, Agent (PG j), trans \ \ set evs \ Says M (PG j) (EncB (priSK M) KM (pubEK (PG j)) AuthReqData P_I) \ set evs" @@ -1151,14 +1151,14 @@ TransStain\; evs \ set_pur; C \ bad; M \ bad|] ==> \OIData OrderDesc KM' trans j' KC' KC'' P_I' P_I''. - HOD = Hash\Number OrderDesc, Number PurchAmt\ & - HOIData = Hash OIData & - Notes M \Number LID_M, Agent (PG j'), trans\ \ set evs & - Says C M \P_I', OIData, Hash PIData\ \ set evs & + HOD = Hash\Number OrderDesc, Number PurchAmt\ \ + HOIData = Hash OIData \ + Notes M \Number LID_M, Agent (PG j'), trans\ \ set evs \ + Says C M \P_I', OIData, Hash PIData\ \ set evs \ Says M (PG j') (EncB (priSK M) KM' (pubEK (PG j')) - AuthReqData P_I'') \ set evs & + AuthReqData P_I'') \ set evs \ P_I' = \PI_sign, - EXcrypt KC' (pubEK (PG j')) \PIHead, Hash OIData\ PANData\ & + EXcrypt KC' (pubEK (PG j')) \PIHead, Hash OIData\ PANData\ \ P_I'' = \PI_sign, EXcrypt KC'' (pubEK (PG j)) \PIHead, Hash OIData\ PANData\" apply clarify diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/SMT_Examples/SMT_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/SMT_Examples/SMT_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -242,7 +242,7 @@ lemma assumes "(\x y. P x y = x)" and "(\x. \y. P x y) = (\x. P x c)" - shows "(EX y. P x y) = P x c" + shows "(\y. P x y) = P x c" using assms by smt lemma @@ -345,8 +345,8 @@ lemma "\x y::int. (2 * x + 1) \ (2 * y)" by smt lemma "\x y::int. x + y > 2 \ x + y = 2 \ x + y < 2" by smt lemma "\x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt -lemma "if (ALL x::int. x < 0 \ x > 0) then False else True" by smt -lemma "(if (ALL x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" by smt +lemma "if (\x::int. x < 0 \ x > 0) then False else True" by smt +lemma "(if (\x::int. x < 0 \ x > 0) then -1 else 3) > (0::int)" by smt lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" by smt lemma "\x::int. \x y. 0 < x \ 0 < y \ (0::int) < x + y" by smt lemma "\u::int. \(x::int) y::real. 0 < x \ 0 < y \ -1 < x" by smt @@ -437,7 +437,7 @@ lemma True using let_rsp by smt lemma "le = (\) \ le (3::int) 42" by smt lemma "map (\i::int. i + 1) [0, 1] = [1, 2]" by (smt list.map) -lemma "(ALL x. P x) \ ~ All P" by smt +lemma "(\x. P x) \ \ All P" by smt fun dec_10 :: "int \ int" where "dec_10 n = (if n < 10 then n else dec_10 (n - 10))" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Set.thy --- a/src/HOL/Set.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Set.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1780,7 +1780,7 @@ for f :: "'a::order \ 'b::order" \ \Courtesy of Stephan Merz\ apply clarify - apply (erule_tac P = "\x. x : S" in LeastI2_order) + apply (erule_tac P = "\x. x \ S" in LeastI2_order) apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Set_Interval.thy Thu Feb 15 13:04:36 2018 +0100 @@ -86,7 +86,7 @@ subsection \Various equivalences\ -lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) -lemma (in ord) atMost_iff [iff]: "(i: atMost k) = (i<=k)" +lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" @@ -1084,8 +1084,7 @@ by (simp add: atLeastAtMost_def) text \A bounded set of natural numbers is finite.\ -lemma bounded_nat_set_is_finite: - "(ALL i:N. i < (n::nat)) ==> finite N" +lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" apply (rule finite_subset) apply (rule_tac [2] finite_lessThan, auto) done @@ -1153,11 +1152,11 @@ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof - show "?A <= ?B" + show "?A \ ?B" proof - fix x assume "x : ?A" - then obtain i where i: "i\n" "x : M i" by auto - show "x : ?B" + fix x assume "x \ ?A" + then obtain i where i: "i\n" "x \ M i" by auto + show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next @@ -1165,18 +1164,18 @@ qed qed next - show "?B <= ?A" by fastforce + show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof - show "?A <= ?B" by fastforce + show "?A \ ?B" by fastforce next - show "?B <= ?A" + show "?B \ ?A" proof - fix x assume "x : ?B" - then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto + fix x assume "x \ ?B" + then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed @@ -1731,7 +1730,7 @@ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well -with the simplifier who adds the unsimplified premise @{term"x:B"} to +with the simplifier who adds the unsimplified premise @{term"x\B"} to the context.\ lemmas sum_ivl_cong = sum.ivl_cong diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/TLA/Memory/Memory.thy --- a/src/HOL/TLA/Memory/Memory.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/TLA/Memory/Memory.thy Thu Feb 15 13:04:36 2018 +0100 @@ -197,7 +197,7 @@ lemma ReadInner_enabled: "\p. basevars (rtrner ch ! p, rs!p) \ \ Calling ch p \ (arg = #(read l)) \ Enabled (ReadInner ch mm rs p l)" - apply (case_tac "l : MemLoc") + apply (case_tac "l \ MemLoc") apply (force simp: ReadInner_def GoodRead_def BadRead_def RdRequest_def intro!: exI elim!: base_enabled [temp_use])+ done @@ -205,7 +205,7 @@ lemma WriteInner_enabled: "\p. basevars (mm!l, rtrner ch ! p, rs!p) \ \ Calling ch p \ (arg = #(write l v)) \ Enabled (WriteInner ch mm rs p l v)" - apply (case_tac "l:MemLoc & v:MemVal") + apply (case_tac "l \ MemLoc \ v \ MemVal") apply (force simp: WriteInner_def GoodWrite_def BadWrite_def WrRequest_def intro!: exI elim!: base_enabled [temp_use])+ done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Thu Feb 15 13:04:36 2018 +0100 @@ -33,15 +33,15 @@ definition (* auxiliary predicates *) MVOKBARF :: "Vals \ bool" - where "MVOKBARF v \ (v : MemVal) | (v = OK) | (v = BadArg) | (v = RPCFailure)" + where "MVOKBARF v \ (v \ MemVal) \ (v = OK) \ (v = BadArg) \ (v = RPCFailure)" definition MVOKBA :: "Vals \ bool" - where "MVOKBA v \ (v : MemVal) | (v = OK) | (v = BadArg)" + where "MVOKBA v \ (v \ MemVal) \ (v = OK) \ (v = BadArg)" definition MVNROKBA :: "Vals \ bool" - where "MVNROKBA v \ (v : MemVal) | (v = NotAResult) | (v = OK) | (v = BadArg)" + where "MVNROKBA v \ (v \ MemVal) \ (v = NotAResult) \ (v = OK) \ (v = BadArg)" definition (* tuples of state functions changed by the various components *) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/TLA/Memory/MemoryParameters.thy --- a/src/HOL/TLA/Memory/MemoryParameters.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/TLA/Memory/MemoryParameters.thy Thu Feb 15 13:04:36 2018 +0100 @@ -29,7 +29,7 @@ (* basic assumptions about the above constants and predicates *) BadArgNoMemVal: "BadArg \ MemVal" and MemFailNoMemVal: "MemFailure \ MemVal" and - InitValMemVal: "InitVal : MemVal" and + InitValMemVal: "InitVal \ MemVal" and NotAResultNotVal: "NotAResult \ MemVal" and NotAResultNotOK: "NotAResult \ OK" and NotAResultNotBA: "NotAResult \ BadArg" and diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML --- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Thu Feb 15 13:04:36 2018 +0100 @@ -264,13 +264,13 @@ | try_dest_All t = t val _ = @{assert} - ((@{term "! x. (! y. P) = True"} + ((@{term "\x. (\y. P) = True"} |> try_dest_All |> Term.strip_abs_vars) = [("x", @{typ "'a"})]) val _ = @{assert} - ((@{prop "! x. (! y. P) = True"} + ((@{prop "\x. (\y. P) = True"} |> try_dest_All |> Term.strip_abs_vars) = [("x", @{typ "'a"})]) @@ -306,7 +306,7 @@ Free ("x", @{typ "'a"}))) in @{assert} - ((@{term "! x. All ((=) x)"} + ((@{term "\x. All ((=) x)"} |> strip_top_All_vars) = answer) end @@ -642,7 +642,7 @@ val simpset = empty_simpset ctxt |> Simplifier.add_simp - @{lemma "! x. P x & Q x \ (! x. P x) & (! x. Q x)" + @{lemma "\x. P x \ Q x \ (\x. P x) \ (\x. Q x)" by (rule eq_reflection, auto)} in CHANGED (asm_full_simp_tac simpset i) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/TPTP/TPTP_Proof_Reconstruction.thy --- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Thu Feb 15 13:04:36 2018 +0100 @@ -346,8 +346,8 @@ "(((A :: bool) = B) = True) == (((A \ B) = True) & ((B \ A) = True))" "((A :: bool) = B) = False == (((~A) | B) = False) | (((~B) | A) = False)" - "((F = G) = True) == (! x. (F x = G x)) = True" - "((F = G) = False) == (! x. (F x = G x)) = False" + "((F = G) = True) == (\x. (F x = G x)) = True" + "((F = G) = False) == (\x. (F x = G x)) = False" "(A | B) = True == (A = True) | (B = True)" "(A & B) = False == (A = False) | (B = False)" @@ -403,7 +403,7 @@ "((A \ B) = (C \ D)) = False \ (B = C) = False \ (A = D) = False" by auto -lemma extuni_func [rule_format]: "(F = G) = False \ (! X. (F X = G X)) = False" by auto +lemma extuni_func [rule_format]: "(F = G) = False \ (\X. (F X = G X)) = False" by auto subsection "Emulation: tactics" @@ -482,18 +482,18 @@ subsection "Skolemisation" lemma skolemise [rule_format]: - "\ P. (~ (! x. P x)) \ ~ (P (SOME x. ~ P x))" + "\ P. (\ (\x. P x)) \ \ (P (SOME x. ~ P x))" proof - - have "\ P. (~ (! x. P x)) \ ~ (P (SOME x. ~ P x))" + have "\ P. (\ (\x. P x)) \ \ (P (SOME x. ~ P x))" proof - fix P - assume ption: "~ (! x. P x)" - hence a: "? x. ~ P x" by force + assume ption: "\ (\x. P x)" + hence a: "\x. \ P x" by force - have hilbert : "\ P. (? x. P x) \ (P (SOME x. P x))" + have hilbert : "\P. (\x. P x) \ (P (SOME x. P x))" proof - fix P - assume "(? x. P x)" + assume "(\x. P x)" thus "(P (SOME x. P x))" apply auto apply (rule someI) @@ -501,9 +501,9 @@ done qed - from a show "~ P (SOME x. ~ P x)" + from a show "\ P (SOME x. \ P x)" proof - - assume "? x. ~ P x" + assume "\x. \ P x" hence "\ P (SOME x. \ P x)" by (rule hilbert) thus ?thesis . qed @@ -512,13 +512,13 @@ qed lemma polar_skolemise [rule_format]: - "\ P. (! x. P x) = False \ (P (SOME x. ~ P x)) = False" + "\P. (\x. P x) = False \ (P (SOME x. \ P x)) = False" proof - - have "\ P. (! x. P x) = False \ (P (SOME x. ~ P x)) = False" + have "\P. (\x. P x) = False \ (P (SOME x. \ P x)) = False" proof - fix P - assume ption: "(! x. P x) = False" - hence "\ (\ x. P x)" by force + assume ption: "(\x. P x) = False" + hence "\ (\x. P x)" by force hence "\ All P" by force hence "\ (P (SOME x. \ P x))" by (rule skolemise) thus "(P (SOME x. \ P x)) = False" by force @@ -527,12 +527,12 @@ qed lemma leo2_skolemise [rule_format]: - "\ P sk. (! x. P x) = False \ (sk = (SOME x. ~ P x)) \ (P sk) = False" + "\P sk. (\x. P x) = False \ (sk = (SOME x. \ P x)) \ (P sk) = False" by (clarify, rule polar_skolemise) lemma lift_forall [rule_format]: - "!! x. (! x. A x) = True ==> (A x) = True" - "!! x. (? x. A x) = False ==> (A x) = False" + "\x. (\x. A x) = True \ (A x) = True" + "\x. (\x. A x) = False \ (A x) = False" by auto lemma lift_exists [rule_format]: "\(All P) = False; sk = (SOME x. \ P x)\ \ P sk = False" @@ -975,7 +975,7 @@ let val simpset = empty_simpset ctxt (*NOTE for some reason, Bind exception gets raised if ctxt's simpset isn't emptied*) - |> Simplifier.add_simp @{lemma "Ex P == (~ (! x. ~ P x))" by auto} + |> Simplifier.add_simp @{lemma "Ex P == (\ (\x. \ P x))" by auto} in CHANGED (asm_full_simp_tac simpset i) end @@ -1541,7 +1541,7 @@ subsubsection "Finite types" (*lift quantification from a singleton literal to a singleton clause*) lemma forall_pos_lift: -"\(! X. P X) = True; ! X. (P X = True) \ R\ \ R" by auto +"\(\X. P X) = True; \X. (P X = True) \ R\ \ R" by auto (*predicate over the type of the leading quantified variable*) @@ -1750,10 +1750,10 @@ \ lemma drop_redundant_literal_qtfr: - "(! X. P) = True \ P = True" - "(? X. P) = True \ P = True" - "(! X. P) = False \ P = False" - "(? X. P) = False \ P = False" + "(\X. P) = True \ P = True" + "(\X. P) = True \ P = True" + "(\X. P) = False \ P = False" + "(\X. P) = False \ P = False" by auto ML \ @@ -1952,11 +1952,11 @@ subsection "Handling split 'preprocessing'" lemma split_tranfs: - "! x. P x & Q x \ (! x. P x) & (! x. Q x)" - "~ (~ A) \ A" - "? x. A \ A" - "(A & B) & C \ A & B & C" - "A = B \ (A --> B) & (B --> A)" + "\x. P x \ Q x \ (\x. P x) \ (\x. Q x)" + "\ (\ A) \ A" + "\x. A \ A" + "(A \ B) \ C \ A \ B \ C" + "A = B \ (A \ B) \ (B \ A)" by (rule eq_reflection, auto)+ (*Same idiom as ex_expander_tac*) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Transitive_Closure.thy Thu Feb 15 13:04:36 2018 +0100 @@ -243,7 +243,7 @@ lemmas rtrancl_converseI = rtranclp_converseI [to_set] -lemma rtrancl_converse: "(r^-1)\<^sup>* = (r\<^sup>*)^-1" +lemma rtrancl_converse: "(r\)\<^sup>* = (r\<^sup>*)\" by (fast dest!: rtrancl_converseD intro!: rtrancl_converseI) lemma sym_rtrancl: "sym r \ sym (r\<^sup>*)" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/Alloc.thy Thu Feb 15 13:04:36 2018 +0100 @@ -79,7 +79,7 @@ definition \ \spec (4)\ client_bounded :: "'a clientState_d program set" - where "client_bounded = UNIV guarantees Always {s. ALL elt : set (ask s). elt \ NbT}" + where "client_bounded = UNIV guarantees Always {s. \elt \ set (ask s). elt \ NbT}" definition \ \spec (5)\ @@ -132,8 +132,8 @@ (INT i : lessThan Nclients. Increasing (sub i o allocAsk) Int Increasing (sub i o allocRel)) Int - Always {s. ALL i NbT} + Always {s. \ielt \ set ((sub i o allocAsk) s). elt \ NbT} Int (INT i : lessThan Nclients. INT h. {s. h \ (sub i o allocGiv)s & h pfixGe (sub i o allocAsk)s} @@ -242,13 +242,13 @@ systemState.dummy = allocState_d.dummy al|))" axiomatization Alloc :: "'a allocState_d program" - where Alloc: "Alloc : alloc_spec" + where Alloc: "Alloc \ alloc_spec" axiomatization Client :: "'a clientState_d program" - where Client: "Client : client_spec" + where Client: "Client \ client_spec" axiomatization Network :: "'a systemState program" - where Network: "Network : network_spec" + where Network: "Network \ network_spec" definition System :: "'a systemState program" where "System = rename sysOfAlloc Alloc \ Network \ @@ -735,22 +735,22 @@ \ text\Lifting \Client_Increasing\ to @{term systemState}\ -lemma rename_Client_Increasing: "i : I - ==> rename sysOfClient (plam x: I. rename client_map Client) : +lemma rename_Client_Increasing: "i \ I + ==> rename sysOfClient (plam x: I. rename client_map Client) \ UNIV guarantees Increasing (ask o sub i o client) Int Increasing (rel o sub i o client)" by rename_client_map -lemma preserves_sub_fst_lift_map: "[| F : preserves w; i ~= j |] - ==> F : preserves (sub i o fst o lift_map j o funPair v w)" +lemma preserves_sub_fst_lift_map: "[| F \ preserves w; i \ j |] + ==> F \ preserves (sub i o fst o lift_map j o funPair v w)" apply (auto simp add: lift_map_def split_def linorder_neq_iff o_def) apply (drule_tac [!] subset_preserves_o [THEN [2] rev_subsetD]) apply (auto simp add: o_def) done lemma client_preserves_giv_oo_client_map: "[| i < Nclients; j < Nclients |] - ==> Client : preserves (giv o sub i o fst o lift_map j o client_map)" + ==> Client \ preserves (giv o sub i o fst o lift_map j o client_map)" apply (cases "i=j") apply (simp, simp add: o_def non_dummy_def) apply (drule Client_preserves_dummy [THEN preserves_sub_fst_lift_map]) @@ -784,7 +784,7 @@ rename_sysOfAlloc_ok_Network [THEN ok_sym] lemma System_Increasing: "i < Nclients - ==> System : Increasing (ask o sub i o client) Int + ==> System \ Increasing (ask o sub i o client) Int Increasing (rel o sub i o client)" apply (rule component_guaranteesD [OF rename_Client_Increasing Client_component_System]) apply auto @@ -803,7 +803,7 @@ allocGiv_o_inv_sysOfAlloc_eq'] lemma System_Increasing_allocGiv: - "i < Nclients ==> System : Increasing (sub i o allocGiv)" + "i < Nclients \ System \ Increasing (sub i o allocGiv)" apply (unfold System_def) apply (simp add: o_def) apply (rule rename_Alloc_Increasing [THEN guarantees_Join_I1, THEN guaranteesD]) @@ -822,19 +822,19 @@ and allows it to be combined using \Always_Int_rule\ below.\ lemma System_Follows_rel: - "i < Nclients ==> System : ((sub i o allocRel) Fols (rel o sub i o client))" + "i < Nclients ==> System \ ((sub i o allocRel) Fols (rel o sub i o client))" apply (auto intro!: Network_Rel [THEN component_guaranteesD]) apply (simp add: ok_commute [of Network]) done lemma System_Follows_ask: - "i < Nclients ==> System : ((sub i o allocAsk) Fols (ask o sub i o client))" + "i < Nclients ==> System \ ((sub i o allocAsk) Fols (ask o sub i o client))" apply (auto intro!: Network_Ask [THEN component_guaranteesD]) apply (simp add: ok_commute [of Network]) done lemma System_Follows_allocGiv: - "i < Nclients ==> System : (giv o sub i o client) Fols (sub i o allocGiv)" + "i < Nclients ==> System \ (giv o sub i o client) Fols (sub i o allocGiv)" apply (auto intro!: Network_Giv [THEN component_guaranteesD] rename_Alloc_Increasing [THEN component_guaranteesD]) apply (simp_all add: o_def non_dummy_def ok_commute [of Network]) @@ -842,21 +842,21 @@ done -lemma Always_giv_le_allocGiv: "System : Always (INT i: lessThan Nclients. +lemma Always_giv_le_allocGiv: "System \ Always (INT i: lessThan Nclients. {s. (giv o sub i o client) s \ (sub i o allocGiv) s})" apply auto apply (erule System_Follows_allocGiv [THEN Follows_Bounded]) done -lemma Always_allocAsk_le_ask: "System : Always (INT i: lessThan Nclients. +lemma Always_allocAsk_le_ask: "System \ Always (INT i: lessThan Nclients. {s. (sub i o allocAsk) s \ (ask o sub i o client) s})" apply auto apply (erule System_Follows_ask [THEN Follows_Bounded]) done -lemma Always_allocRel_le_rel: "System : Always (INT i: lessThan Nclients. +lemma Always_allocRel_le_rel: "System \ Always (INT i: lessThan Nclients. {s. (sub i o allocRel) s \ (rel o sub i o client) s})" by (auto intro!: Follows_Bounded System_Follows_rel) @@ -875,7 +875,7 @@ text\safety (1), step 3\ lemma System_sum_bounded: - "System : Always {s. (\i \ lessThan Nclients. (tokens o sub i o allocGiv) s) + "System \ Always {s. (\i \ lessThan Nclients. (tokens o sub i o allocGiv) s) \ NbT + (\i \ lessThan Nclients. (tokens o sub i o allocRel) s)}" apply (simp add: o_apply) apply (insert Alloc_Safety [THEN rename_guarantees_sysOfAlloc_I]) @@ -886,14 +886,14 @@ text\Follows reasoning\ -lemma Always_tokens_giv_le_allocGiv: "System : Always (INT i: lessThan Nclients. +lemma Always_tokens_giv_le_allocGiv: "System \ Always (INT i: lessThan Nclients. {s. (tokens o giv o sub i o client) s \ (tokens o sub i o allocGiv) s})" apply (rule Always_giv_le_allocGiv [THEN Always_weaken]) apply (auto intro: tokens_mono_prefix simp add: o_apply) done -lemma Always_tokens_allocRel_le_rel: "System : Always (INT i: lessThan Nclients. +lemma Always_tokens_allocRel_le_rel: "System \ Always (INT i: lessThan Nclients. {s. (tokens o sub i o allocRel) s \ (tokens o rel o sub i o client) s})" apply (rule Always_allocRel_le_rel [THEN Always_weaken]) @@ -901,7 +901,7 @@ done text\safety (1), step 4 (final result!)\ -theorem System_safety: "System : system_safety" +theorem System_safety: "System \ system_safety" apply (unfold system_safety_def) apply (tactic \resolve_tac @{context} [Always_Int_rule [@{thm System_sum_bounded}, @{thm Always_tokens_giv_le_allocGiv}, @{thm Always_tokens_allocRel_le_rel}] RS @@ -924,27 +924,27 @@ lemmas System_Increasing_allocAsk = System_Follows_ask [THEN Follows_Increasing1] text\progress (2), step 3: lifting \Client_Bounded\ to systemState\ -lemma rename_Client_Bounded: "i : I - ==> rename sysOfClient (plam x: I. rename client_map Client) : +lemma rename_Client_Bounded: "i \ I + ==> rename sysOfClient (plam x: I. rename client_map Client) \ UNIV guarantees - Always {s. ALL elt : set ((ask o sub i o client) s). elt \ NbT}" + Always {s. \elt \ set ((ask o sub i o client) s). elt \ NbT}" by rename_client_map lemma System_Bounded_ask: "i < Nclients - ==> System : Always - {s. ALL elt : set ((ask o sub i o client) s). elt \ NbT}" + ==> System \ Always + {s. \elt \ set ((ask o sub i o client) s). elt \ NbT}" apply (rule component_guaranteesD [OF rename_Client_Bounded Client_component_System]) apply auto done -lemma Collect_all_imp_eq: "{x. ALL y. P y --> Q x y} = (INT y: {y. P y}. {x. Q x y})" +lemma Collect_all_imp_eq: "{x. \y. P y \ Q x y} = (INT y: {y. P y}. {x. Q x y})" apply blast done text\progress (2), step 4\ -lemma System_Bounded_allocAsk: "System : Always {s. ALL i NbT}" +lemma System_Bounded_allocAsk: "System \ Always {s. ALL ielt \ set ((sub i o allocAsk) s). elt \ NbT}" apply (auto simp add: Collect_all_imp_eq) apply (tactic \resolve_tac @{context} [Always_Int_rule [@{thm Always_allocAsk_le_ask}, @{thm System_Bounded_ask}] RS @{thm Always_weaken}] 1\) @@ -958,9 +958,9 @@ lemmas System_Increasing_giv = System_Follows_allocGiv [THEN Follows_Increasing1] -lemma rename_Client_Progress: "i: I +lemma rename_Client_Progress: "i \ I ==> rename sysOfClient (plam x: I. rename client_map Client) - : Increasing (giv o sub i o client) + \ Increasing (giv o sub i o client) guarantees (INT h. {s. h \ (giv o sub i o client) s & h pfixGe (ask o sub i o client) s} @@ -972,7 +972,7 @@ text\progress (2), step 7\ lemma System_Client_Progress: - "System : (INT i : (lessThan Nclients). + "System \ (INT i : (lessThan Nclients). INT h. {s. h \ (giv o sub i o client) s & h pfixGe (ask o sub i o client) s} LeadsTo {s. tokens h \ (tokens o rel o sub i o client) s})" @@ -998,7 +998,7 @@ lemma System_lemma3: "i < Nclients - ==> System : {s. h \ (sub i o allocGiv) s & + ==> System \ {s. h \ (sub i o allocGiv) s & h pfixGe (sub i o allocAsk) s} LeadsTo {s. h \ (giv o sub i o client) s & @@ -1013,7 +1013,7 @@ text\progress (2), step 8: Client i's "release" action is visible system-wide\ lemma System_Alloc_Client_Progress: "i < Nclients - ==> System : {s. h \ (sub i o allocGiv) s & + ==> System \ {s. h \ (sub i o allocGiv) s & h pfixGe (sub i o allocAsk) s} LeadsTo {s. tokens h \ (tokens o sub i o allocRel) s}" apply (rule LeadsTo_Trans) @@ -1033,7 +1033,7 @@ text\progress (2), step 9\ lemma System_Alloc_Progress: - "System : (INT i : (lessThan Nclients). + "System \ (INT i : (lessThan Nclients). INT h. {s. h \ (sub i o allocAsk) s} LeadsTo {s. h pfixLe (sub i o allocGiv) s})" apply (simp only: o_apply sub_def) @@ -1048,7 +1048,7 @@ done text\progress (2), step 10 (final result!)\ -lemma System_Progress: "System : system_progress" +lemma System_Progress: "System \ system_progress" apply (unfold system_progress_def) apply (cut_tac System_Alloc_Progress) apply auto @@ -1058,7 +1058,7 @@ done -theorem System_correct: "System : system_spec" +theorem System_correct: "System \ system_spec" apply (unfold system_spec_def) apply (blast intro: System_safety System_Progress) done @@ -1083,7 +1083,7 @@ done text\Could go to Extend.ML\ -lemma bij_fst_inv_inv_eq: "bij f ==> fst (inv (%(x, u). inv f x) z) = f z" +lemma bij_fst_inv_inv_eq: "bij f \ fst (inv (%(x, u). inv f x) z) = f z" apply (rule fst_inv_equalityI) apply (rule_tac f = "%z. (f z, h z)" for h in surjI) apply (simp add: bij_is_inj inv_f_f) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/AllocBase.thy --- a/src/HOL/UNITY/Comp/AllocBase.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/AllocBase.thy Thu Feb 15 13:04:36 2018 +0100 @@ -76,7 +76,7 @@ by (simp add: bag_of_nths_Un_Int [symmetric]) lemma bag_of_nths_UN_disjoint [rule_format]: - "[| finite I; ALL i:I. ALL j:I. i~=j --> A i Int A j = {} |] + "[| finite I; \i\I. \j\I. i\j \ A i Int A j = {} |] ==> bag_of (nths l (UNION I A)) = (\i\I. bag_of (nths l (A i)))" apply (auto simp add: bag_of_nths) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/Handshake.thy --- a/src/HOL/UNITY/Comp/Handshake.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/Handshake.thy Thu Feb 15 13:04:36 2018 +0100 @@ -47,7 +47,7 @@ invFG_def [THEN def_set_simp, simp] -lemma invFG: "(F \ G) : Always invFG" +lemma invFG: "(F \ G) \ Always invFG" apply (rule AlwaysI) apply force apply (rule constrains_imp_Constrains [THEN StableI]) @@ -56,21 +56,21 @@ apply (unfold G_def, safety) done -lemma lemma2_1: "(F \ G) : ({s. NF s = k} - {s. BB s}) LeadsTo +lemma lemma2_1: "(F \ G) \ ({s. NF s = k} - {s. BB s}) LeadsTo ({s. NF s = k} Int {s. BB s})" apply (rule stable_Join_ensures1[THEN leadsTo_Basis, THEN leadsTo_imp_LeadsTo]) apply (unfold F_def, safety) apply (unfold G_def, ensures_tac "cmdG") done -lemma lemma2_2: "(F \ G) : ({s. NF s = k} Int {s. BB s}) LeadsTo +lemma lemma2_2: "(F \ G) \ ({s. NF s = k} Int {s. BB s}) LeadsTo {s. k < NF s}" apply (rule stable_Join_ensures2[THEN leadsTo_Basis, THEN leadsTo_imp_LeadsTo]) apply (unfold F_def, ensures_tac "cmdF") apply (unfold G_def, safety) done -lemma progress: "(F \ G) : UNIV LeadsTo {s. m < NF s}" +lemma progress: "(F \ G) \ UNIV LeadsTo {s. m < NF s}" apply (rule LeadsTo_weaken_R) apply (rule_tac f = "NF" and l = "Suc m" and B = "{}" in GreaterThan_bounded_induct) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/Priority.thy --- a/src/HOL/UNITY/Comp/Priority.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/Priority.thy Thu Feb 15 13:04:36 2018 +0100 @@ -84,13 +84,13 @@ by (simp add: Component_def, safety, auto) text\property 4\ -lemma Component_waits_priority: "Component i: {s. ((i,j):s) = b} Int (- Highest i) co {s. ((i,j):s)=b}" +lemma Component_waits_priority: "Component i \ {s. ((i,j) \ s) = b} \ (- Highest i) co {s. ((i,j) \ s)=b}" by (simp add: Component_def, safety) text\property 5: charpentier and Chandy mistakenly express it as 'transient Highest i'. Consider the case where i has neighbors\ lemma Component_yields_priority: - "Component i: {s. neighbors i s \ {}} Int Highest i + "Component i \ {s. neighbors i s \ {}} Int Highest i ensures - Highest i" apply (simp add: Component_def) apply (ensures_tac "act i", blast+) @@ -107,7 +107,7 @@ by (simp add: Component_def, safety, fast) text\property 7: local axiom\ -lemma locality: "Component i \ stable {s. \j k. j\i & k\i--> ((j,k):s) = b j k}" +lemma locality: "Component i \ stable {s. \j k. j\i & k\i--> ((j,k) \ s) = b j k}" by (simp add: Component_def, safety) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/PriorityAux.thy --- a/src/HOL/UNITY/Comp/PriorityAux.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/PriorityAux.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,33 +12,33 @@ typedecl vertex definition symcl :: "(vertex*vertex)set=>(vertex*vertex)set" where - "symcl r == r \ (r^-1)" + "symcl r == r \ (r\)" \ \symmetric closure: removes the orientation of a relation\ definition neighbors :: "[vertex, (vertex*vertex)set]=>vertex set" where - "neighbors i r == ((r \ r^-1)``{i}) - {i}" + "neighbors i r == ((r \ r\)``{i}) - {i}" \ \Neighbors of a vertex i\ definition R :: "[vertex, (vertex*vertex)set]=>vertex set" where "R i r == r``{i}" definition A :: "[vertex, (vertex*vertex)set]=>vertex set" where - "A i r == (r^-1)``{i}" + "A i r == (r\)``{i}" definition reach :: "[vertex, (vertex*vertex)set]=> vertex set" where - "reach i r == (r^+)``{i}" + "reach i r == (r\<^sup>+)``{i}" \ \reachable and above vertices: the original notation was R* and A*\ definition above :: "[vertex, (vertex*vertex)set]=> vertex set" where - "above i r == ((r^-1)^+)``{i}" + "above i r == ((r\)\<^sup>+)``{i}" definition reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set" where - "reverse i r == (r - {(x,y). x=i | y=i} \ r) \ ({(x,y). x=i|y=i} \ r)^-1" + "reverse i r == (r - {(x,y). x=i | y=i} \ r) \ ({(x,y). x=i|y=i} \ r)\" definition derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where \ \The original definition\ "derive1 i r q == symcl r = symcl q & - (\k k'. k\i & k'\i -->((k,k'):r) = ((k,k'):q)) & + (\k k'. k\i & k'\i -->((k,k') \ r) = ((k,k') \ q)) \ A i r = {} & R i q = {}" definition derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where @@ -68,13 +68,13 @@ (* The equalities (above i r = {}) = (A i r = {}) and (reach i r = {}) = (R i r) rely on the following theorem *) -lemma image0_trancl_iff_image0_r: "((r^+)``{i} = {}) = (r``{i} = {})" +lemma image0_trancl_iff_image0_r: "((r\<^sup>+)``{i} = {}) = (r``{i} = {})" apply auto apply (erule trancl_induct, auto) done (* Another form usefull in some situation *) -lemma image0_r_iff_image0_trancl: "(r``{i}={}) = (ALL x. ((i,x):r^+) = False)" +lemma image0_r_iff_image0_trancl: "(r``{i}={}) = (\x. ((i,x) \ r\<^sup>+) = False)" apply auto apply (drule image0_trancl_iff_image0_r [THEN ssubst], auto) done @@ -117,7 +117,7 @@ (* Lemma 2 *) lemma maximal_converse_image0: - "(z, i):r^+ ==> (\y. (y, z):r --> (y,i) \ r^+) = ((r^-1)``{z}={})" + "(z, i) \ r\<^sup>+ \ (\y. (y, z) \ r \ (y,i) \ r\<^sup>+) = ((r\)``{z}={})" apply auto apply (frule_tac r = r in trancl_into_trancl2, auto) done @@ -125,7 +125,7 @@ lemma above_lemma_a: "acyclic r ==> A i r\{}-->(\j \ above i r. A j r = {})" apply (simp add: acyclic_eq_wf wf_eq_minimal) -apply (drule_tac x = " ((r^-1) ^+) ``{i}" in spec) +apply (drule_tac x = " ((r\)\<^sup>+) ``{i}" in spec) apply auto apply (simp add: maximal_converse_image0 trancl_converse) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Comp/TimerArray.thy --- a/src/HOL/UNITY/Comp/TimerArray.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Comp/TimerArray.thy Thu Feb 15 13:04:36 2018 +0100 @@ -24,14 +24,14 @@ declare count_def [simp] decr_def [simp] (*Demonstrates induction, but not used in the following proof*) -lemma Timer_leadsTo_zero: "Timer : UNIV leadsTo {s. count s = 0}" +lemma Timer_leadsTo_zero: "Timer \ UNIV leadsTo {s. count s = 0}" apply (rule_tac f = count in lessThan_induct, simp) apply (case_tac "m") apply (force intro!: subset_imp_leadsTo) apply (unfold Timer_def, ensures_tac "decr") done -lemma Timer_preserves_snd [iff]: "Timer : preserves snd" +lemma Timer_preserves_snd [iff]: "Timer \ preserves snd" apply (rule preservesI) apply (unfold Timer_def, safety) done @@ -41,8 +41,8 @@ lemma TimerArray_leadsTo_zero: "finite I - ==> (plam i: I. Timer) : UNIV leadsTo {(s,uu). ALL i:I. s i = 0}" -apply (erule_tac A'1 = "%i. lift_set i ({0} \ UNIV)" + \ (plam i: I. Timer) \ UNIV leadsTo {(s,uu). \i\I. s i = 0}" +apply (erule_tac A'1 = "\i. lift_set i ({0} \ UNIV)" in finite_stable_completion [THEN leadsTo_weaken]) apply auto (*Safety property, already reduced to the single Timer case*) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Constrains.thy --- a/src/HOL/UNITY/Constrains.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Constrains.thy Thu Feb 15 13:04:36 2018 +0100 @@ -19,7 +19,7 @@ (*Initial trace is empty*) Init: "s \ init ==> (s,[]) \ traces init acts" - | Acts: "[| act: acts; (s,evs) \ traces init acts; (s,s'): act |] + | Acts: "[| act \ acts; (s,evs) \ traces init acts; (s,s') \ act |] ==> (s', s#evs) \ traces init acts" @@ -29,7 +29,7 @@ where Init: "s \ Init F ==> s \ reachable F" - | Acts: "[| act: Acts F; s \ reachable F; (s,s'): act |] + | Acts: "[| act \ Acts F; s \ reachable F; (s,s') \ act |] ==> s' \ reachable F" definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where @@ -106,7 +106,7 @@ done lemma ConstrainsI: - "(!!act s s'. [| act: Acts F; (s,s') \ act; s \ A |] ==> s': A') + "(!!act s s'. [| act \ Acts F; (s,s') \ act; s \ A |] ==> s' \ A') ==> F \ A Co A'" apply (rule constrains_imp_Constrains) apply (blast intro: constrainsI) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/ELT.thy --- a/src/HOL/UNITY/ELT.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/ELT.thy Thu Feb 15 13:04:36 2018 +0100 @@ -31,11 +31,11 @@ for CC :: "'a set set" and F :: "'a program" where - Basis: "[| F : A ensures B; A-B : (insert {} CC) |] ==> (A,B) : elt CC F" + Basis: "[| F \ A ensures B; A-B \ (insert {} CC) |] ==> (A,B) \ elt CC F" - | Trans: "[| (A,B) : elt CC F; (B,C) : elt CC F |] ==> (A,C) : elt CC F" + | Trans: "[| (A,B) \ elt CC F; (B,C) \ elt CC F |] ==> (A,C) \ elt CC F" - | Union: "ALL A: S. (A,B) : elt CC F ==> (Union S, B) : elt CC F" + | Union: "\A\S. (A,B) \ elt CC F ==> (Union S, B) \ elt CC F" definition @@ -47,13 +47,13 @@ (*visible version of the LEADS-TO relation*) leadsETo :: "['a set, 'a set set, 'a set] => 'a program set" ("(3_/ leadsTo[_]/ _)" [80,0,80] 80) - where "leadsETo A CC B = {F. (A,B) : elt CC F}" + where "leadsETo A CC B = {F. (A,B) \ elt CC F}" definition LeadsETo :: "['a set, 'a set set, 'a set] => 'a program set" ("(3_/ LeadsTo[_]/ _)" [80,0,80] 80) where "LeadsETo A CC B = - {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] B}" + {F. F \ (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] B}" (*** givenBy ***) @@ -61,35 +61,35 @@ lemma givenBy_id [simp]: "givenBy id = UNIV" by (unfold givenBy_def, auto) -lemma givenBy_eq_all: "(givenBy v) = {A. ALL x:A. ALL y. v x = v y --> y: A}" +lemma givenBy_eq_all: "(givenBy v) = {A. \x\A. \y. v x = v y \ y \ A}" apply (unfold givenBy_def, safe) apply (rule_tac [2] x = "v ` _" in image_eqI, auto) done -lemma givenByI: "(!!x y. [| x:A; v x = v y |] ==> y: A) ==> A: givenBy v" +lemma givenByI: "(\x y. [| x \ A; v x = v y |] ==> y \ A) ==> A \ givenBy v" by (subst givenBy_eq_all, blast) -lemma givenByD: "[| A: givenBy v; x:A; v x = v y |] ==> y: A" +lemma givenByD: "[| A \ givenBy v; x \ A; v x = v y |] ==> y \ A" by (unfold givenBy_def, auto) -lemma empty_mem_givenBy [iff]: "{} : givenBy v" +lemma empty_mem_givenBy [iff]: "{} \ givenBy v" by (blast intro!: givenByI) -lemma givenBy_imp_eq_Collect: "A: givenBy v ==> EX P. A = {s. P(v s)}" -apply (rule_tac x = "%n. EX s. v s = n & s : A" in exI) +lemma givenBy_imp_eq_Collect: "A \ givenBy v ==> \P. A = {s. P(v s)}" +apply (rule_tac x = "\n. \s. v s = n \ s \ A" in exI) apply (simp (no_asm_use) add: givenBy_eq_all) apply blast done -lemma Collect_mem_givenBy: "{s. P(v s)} : givenBy v" +lemma Collect_mem_givenBy: "{s. P(v s)} \ givenBy v" by (unfold givenBy_def, best) -lemma givenBy_eq_Collect: "givenBy v = {A. EX P. A = {s. P(v s)}}" +lemma givenBy_eq_Collect: "givenBy v = {A. \P. A = {s. P(v s)}}" by (blast intro: Collect_mem_givenBy givenBy_imp_eq_Collect) (*preserving v preserves properties given by v*) lemma preserves_givenBy_imp_stable: - "[| F : preserves v; D : givenBy v |] ==> F : stable D" + "[| F \ preserves v; D \ givenBy v |] ==> F \ stable D" by (force simp add: preserves_subset_stable [THEN subsetD] givenBy_eq_Collect) lemma givenBy_o_subset: "givenBy (w o v) <= givenBy v" @@ -98,7 +98,7 @@ done lemma givenBy_DiffI: - "[| A : givenBy v; B : givenBy v |] ==> A-B : givenBy v" + "[| A \ givenBy v; B \ givenBy v |] ==> A-B \ givenBy v" apply (simp (no_asm_use) add: givenBy_eq_Collect) apply safe apply (rule_tac x = "%z. R z & ~ Q z" for R Q in exI) @@ -110,13 +110,13 @@ (** Standard leadsTo rules **) lemma leadsETo_Basis [intro]: - "[| F: A ensures B; A-B: insert {} CC |] ==> F : A leadsTo[CC] B" + "[| F \ A ensures B; A-B \ insert {} CC |] ==> F \ A leadsTo[CC] B" apply (unfold leadsETo_def) apply (blast intro: elt.Basis) done lemma leadsETo_Trans: - "[| F : A leadsTo[CC] B; F : B leadsTo[CC] C |] ==> F : A leadsTo[CC] C" + "[| F \ A leadsTo[CC] B; F \ B leadsTo[CC] C |] ==> F \ A leadsTo[CC] C" apply (unfold leadsETo_def) apply (blast intro: elt.Trans) done @@ -124,33 +124,33 @@ (*Useful with cancellation, disjunction*) lemma leadsETo_Un_duplicate: - "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'" + "F \ A leadsTo[CC] (A' \ A') \ F \ A leadsTo[CC] A'" by (simp add: Un_ac) lemma leadsETo_Un_duplicate2: - "F : A leadsTo[CC] (A' Un C Un C) ==> F : A leadsTo[CC] (A' Un C)" + "F \ A leadsTo[CC] (A' \ C \ C) ==> F \ A leadsTo[CC] (A' Un C)" by (simp add: Un_ac) (*The Union introduction rule as we should have liked to state it*) lemma leadsETo_Union: - "(!!A. A : S ==> F : A leadsTo[CC] B) ==> F : (\S) leadsTo[CC] B" + "(\A. A \ S \ F \ A leadsTo[CC] B) \ F \ (\S) leadsTo[CC] B" apply (unfold leadsETo_def) apply (blast intro: elt.Union) done lemma leadsETo_UN: - "(!!i. i : I ==> F : (A i) leadsTo[CC] B) - ==> F : (UN i:I. A i) leadsTo[CC] B" + "(\i. i \ I \ F \ (A i) leadsTo[CC] B) + ==> F \ (UN i:I. A i) leadsTo[CC] B" apply (blast intro: leadsETo_Union) done (*The INDUCTION rule as we should have liked to state it*) lemma leadsETo_induct: - "[| F : za leadsTo[CC] zb; - !!A B. [| F : A ensures B; A-B : insert {} CC |] ==> P A B; - !!A B C. [| F : A leadsTo[CC] B; P A B; F : B leadsTo[CC] C; P B C |] + "[| F \ za leadsTo[CC] zb; + !!A B. [| F \ A ensures B; A-B \ insert {} CC |] ==> P A B; + !!A B C. [| F \ A leadsTo[CC] B; P A B; F \ B leadsTo[CC] C; P B C |] ==> P A C; - !!B S. ALL A:S. F : A leadsTo[CC] B & P A B ==> P (\S) B + !!B S. \A\S. F \ A leadsTo[CC] B & P A B ==> P (\S) B |] ==> P za zb" apply (unfold leadsETo_def) apply (drule CollectD) @@ -169,13 +169,13 @@ done lemma leadsETo_Trans_Un: - "[| F : A leadsTo[CC] B; F : B leadsTo[DD] C |] - ==> F : A leadsTo[CC Un DD] C" + "[| F \ A leadsTo[CC] B; F \ B leadsTo[DD] C |] + ==> F \ A leadsTo[CC Un DD] C" by (blast intro: leadsETo_mono [THEN subsetD] leadsETo_Trans) lemma leadsETo_Union_Int: - "(!!A. A : S ==> F : (A Int C) leadsTo[CC] B) - ==> F : (\S Int C) leadsTo[CC] B" + "(!!A. A \ S ==> F \ (A Int C) leadsTo[CC] B) + ==> F \ (\S Int C) leadsTo[CC] B" apply (unfold leadsETo_def) apply (simp only: Int_Union_Union) apply (blast intro: elt.Union) @@ -183,16 +183,16 @@ (*Binary union introduction rule*) lemma leadsETo_Un: - "[| F : A leadsTo[CC] C; F : B leadsTo[CC] C |] - ==> F : (A Un B) leadsTo[CC] C" + "[| F \ A leadsTo[CC] C; F \ B leadsTo[CC] C |] + ==> F \ (A Un B) leadsTo[CC] C" using leadsETo_Union [of "{A, B}" F CC C] by auto lemma single_leadsETo_I: - "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B" + "(\x. x \ A ==> F \ {x} leadsTo[CC] B) \ F \ A leadsTo[CC] B" by (subst UN_singleton [symmetric], rule leadsETo_UN, blast) -lemma subset_imp_leadsETo: "A<=B ==> F : A leadsTo[CC] B" +lemma subset_imp_leadsETo: "A<=B \ F \ A leadsTo[CC] B" by (simp add: subset_imp_ensures [THEN leadsETo_Basis] Diff_eq_empty_iff [THEN iffD2]) @@ -203,73 +203,73 @@ (** Weakening laws **) lemma leadsETo_weaken_R: - "[| F : A leadsTo[CC] A'; A'<=B' |] ==> F : A leadsTo[CC] B'" + "[| F \ A leadsTo[CC] A'; A'<=B' |] ==> F \ A leadsTo[CC] B'" by (blast intro: subset_imp_leadsETo leadsETo_Trans) lemma leadsETo_weaken_L: - "[| F : A leadsTo[CC] A'; B<=A |] ==> F : B leadsTo[CC] A'" + "[| F \ A leadsTo[CC] A'; B<=A |] ==> F \ B leadsTo[CC] A'" by (blast intro: leadsETo_Trans subset_imp_leadsETo) (*Distributes over binary unions*) lemma leadsETo_Un_distrib: - "F : (A Un B) leadsTo[CC] C = - (F : A leadsTo[CC] C & F : B leadsTo[CC] C)" + "F \ (A Un B) leadsTo[CC] C = + (F \ A leadsTo[CC] C \ F \ B leadsTo[CC] C)" by (blast intro: leadsETo_Un leadsETo_weaken_L) lemma leadsETo_UN_distrib: - "F : (UN i:I. A i) leadsTo[CC] B = - (ALL i : I. F : (A i) leadsTo[CC] B)" + "F \ (UN i:I. A i) leadsTo[CC] B = + (\i\I. F \ (A i) leadsTo[CC] B)" by (blast intro: leadsETo_UN leadsETo_weaken_L) lemma leadsETo_Union_distrib: - "F : (\S) leadsTo[CC] B = (ALL A : S. F : A leadsTo[CC] B)" + "F \ (\S) leadsTo[CC] B = (\A\S. F \ A leadsTo[CC] B)" by (blast intro: leadsETo_Union leadsETo_weaken_L) lemma leadsETo_weaken: - "[| F : A leadsTo[CC'] A'; B<=A; A'<=B'; CC' <= CC |] - ==> F : B leadsTo[CC] B'" + "[| F \ A leadsTo[CC'] A'; B<=A; A'<=B'; CC' <= CC |] + ==> F \ B leadsTo[CC] B'" apply (drule leadsETo_mono [THEN subsetD], assumption) apply (blast del: subsetCE intro: leadsETo_weaken_R leadsETo_weaken_L leadsETo_Trans) done lemma leadsETo_givenBy: - "[| F : A leadsTo[CC] A'; CC <= givenBy v |] - ==> F : A leadsTo[givenBy v] A'" + "[| F \ A leadsTo[CC] A'; CC <= givenBy v |] + ==> F \ A leadsTo[givenBy v] A'" by (blast intro: leadsETo_weaken) (*Set difference*) lemma leadsETo_Diff: - "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C |] - ==> F : A leadsTo[CC] C" + "[| F \ (A-B) leadsTo[CC] C; F \ B leadsTo[CC] C |] + ==> F \ A leadsTo[CC] C" by (blast intro: leadsETo_Un leadsETo_weaken) (*Binary union version*) lemma leadsETo_Un_Un: - "[| F : A leadsTo[CC] A'; F : B leadsTo[CC] B' |] - ==> F : (A Un B) leadsTo[CC] (A' Un B')" + "[| F \ A leadsTo[CC] A'; F \ B leadsTo[CC] B' |] + ==> F \ (A Un B) leadsTo[CC] (A' Un B')" by (blast intro: leadsETo_Un leadsETo_weaken_R) (** The cancellation law **) lemma leadsETo_cancel2: - "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B' |] - ==> F : A leadsTo[CC] (A' Un B')" + "[| F \ A leadsTo[CC] (A' Un B); F \ B leadsTo[CC] B' |] + ==> F \ A leadsTo[CC] (A' Un B')" by (blast intro: leadsETo_Un_Un subset_imp_leadsETo leadsETo_Trans) lemma leadsETo_cancel1: - "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B' |] - ==> F : A leadsTo[CC] (B' Un A')" + "[| F \ A leadsTo[CC] (B Un A'); F \ B leadsTo[CC] B' |] + ==> F \ A leadsTo[CC] (B' Un A')" apply (simp add: Un_commute) apply (blast intro!: leadsETo_cancel2) done lemma leadsETo_cancel_Diff1: - "[| F : A leadsTo[CC] (B Un A'); F : (B-A') leadsTo[CC] B' |] - ==> F : A leadsTo[CC] (B' Un A')" + "[| F \ A leadsTo[CC] (B Un A'); F \ (B-A') leadsTo[CC] B' |] + ==> F \ A leadsTo[CC] (B' Un A')" apply (rule leadsETo_cancel1) prefer 2 apply assumption apply simp_all @@ -280,8 +280,8 @@ (*Special case of PSP: Misra's "stable conjunction"*) lemma e_psp_stable: - "[| F : A leadsTo[CC] A'; F : stable B; ALL C:CC. C Int B : CC |] - ==> F : (A Int B) leadsTo[CC] (A' Int B)" + "[| F \ A leadsTo[CC] A'; F \ stable B; \C\CC. C Int B \ CC |] + ==> F \ (A Int B) leadsTo[CC] (A' Int B)" apply (unfold stable_def) apply (erule leadsETo_induct) prefer 3 apply (blast intro: leadsETo_Union_Int) @@ -294,14 +294,14 @@ done lemma e_psp_stable2: - "[| F : A leadsTo[CC] A'; F : stable B; ALL C:CC. C Int B : CC |] - ==> F : (B Int A) leadsTo[CC] (B Int A')" + "[| F \ A leadsTo[CC] A'; F \ stable B; \C\CC. C Int B \ CC |] + ==> F \ (B Int A) leadsTo[CC] (B Int A')" by (simp (no_asm_simp) add: e_psp_stable Int_ac) lemma e_psp: - "[| F : A leadsTo[CC] A'; F : B co B'; - ALL C:CC. C Int B Int B' : CC |] - ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))" + "[| F \ A leadsTo[CC] A'; F \ B co B'; + \C\CC. C Int B Int B' \ CC |] + ==> F \ (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))" apply (erule leadsETo_induct) prefer 3 apply (blast intro: leadsETo_Union_Int) (*Transitivity case has a delicate argument involving "cancellation"*) @@ -318,9 +318,9 @@ done lemma e_psp2: - "[| F : A leadsTo[CC] A'; F : B co B'; - ALL C:CC. C Int B Int B' : CC |] - ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))" + "[| F \ A leadsTo[CC] A'; F \ B co B'; + \C\CC. C Int B Int B' \ CC |] + ==> F \ (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))" by (simp add: e_psp Int_ac) @@ -328,9 +328,9 @@ (*??IS THIS NEEDED?? or is it just an example of what's provable??*) lemma gen_leadsETo_imp_Join_leadsETo: - "[| F: (A leadsTo[givenBy v] B); G : preserves v; - F\G : stable C |] - ==> F\G : ((C Int A) leadsTo[(%D. C Int D) ` givenBy v] B)" + "[| F \ (A leadsTo[givenBy v] B); G \ preserves v; + F\G \ stable C |] + ==> F\G \ ((C Int A) leadsTo[(%D. C Int D) ` givenBy v] B)" apply (erule leadsETo_induct) prefer 3 apply (subst Int_Union) @@ -371,7 +371,7 @@ lemma LeadsETo_eq_leadsETo: "A LeadsTo[CC] B = - {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] + {F. F \ (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] (reachable F Int B)}" apply (unfold LeadsETo_def) apply (blast dest: e_psp_stable2 intro: leadsETo_weaken) @@ -380,38 +380,38 @@ (*** Introduction rules: Basis, Trans, Union ***) lemma LeadsETo_Trans: - "[| F : A LeadsTo[CC] B; F : B LeadsTo[CC] C |] - ==> F : A LeadsTo[CC] C" + "[| F \ A LeadsTo[CC] B; F \ B LeadsTo[CC] C |] + ==> F \ A LeadsTo[CC] C" apply (simp add: LeadsETo_eq_leadsETo) apply (blast intro: leadsETo_Trans) done lemma LeadsETo_Union: - "(!!A. A : S ==> F : A LeadsTo[CC] B) ==> F : (\S) LeadsTo[CC] B" + "(\A. A \ S \ F \ A LeadsTo[CC] B) \ F \ (\S) LeadsTo[CC] B" apply (simp add: LeadsETo_def) apply (subst Int_Union) apply (blast intro: leadsETo_UN) done lemma LeadsETo_UN: - "(!!i. i : I ==> F : (A i) LeadsTo[CC] B) - ==> F : (UN i:I. A i) LeadsTo[CC] B" + "(\i. i \ I \ F \ (A i) LeadsTo[CC] B) + \ F \ (UN i:I. A i) LeadsTo[CC] B" apply (blast intro: LeadsETo_Union) done (*Binary union introduction rule*) lemma LeadsETo_Un: - "[| F : A LeadsTo[CC] C; F : B LeadsTo[CC] C |] - ==> F : (A Un B) LeadsTo[CC] C" + "[| F \ A LeadsTo[CC] C; F \ B LeadsTo[CC] C |] + ==> F \ (A Un B) LeadsTo[CC] C" using LeadsETo_Union [of "{A, B}" F CC C] by auto (*Lets us look at the starting state*) lemma single_LeadsETo_I: - "(!!s. s : A ==> F : {s} LeadsTo[CC] B) ==> F : A LeadsTo[CC] B" + "(\s. s \ A ==> F \ {s} LeadsTo[CC] B) \ F \ A LeadsTo[CC] B" by (subst UN_singleton [symmetric], rule LeadsETo_UN, blast) lemma subset_imp_LeadsETo: - "A <= B ==> F : A LeadsTo[CC] B" + "A <= B \ F \ A LeadsTo[CC] B" apply (simp (no_asm) add: LeadsETo_def) apply (blast intro: subset_imp_leadsETo) done @@ -419,21 +419,21 @@ lemmas empty_LeadsETo = empty_subsetI [THEN subset_imp_LeadsETo] lemma LeadsETo_weaken_R: - "[| F : A LeadsTo[CC] A'; A' <= B' |] ==> F : A LeadsTo[CC] B'" + "[| F \ A LeadsTo[CC] A'; A' <= B' |] ==> F \ A LeadsTo[CC] B'" apply (simp add: LeadsETo_def) apply (blast intro: leadsETo_weaken_R) done lemma LeadsETo_weaken_L: - "[| F : A LeadsTo[CC] A'; B <= A |] ==> F : B LeadsTo[CC] A'" + "[| F \ A LeadsTo[CC] A'; B <= A |] ==> F \ B LeadsTo[CC] A'" apply (simp add: LeadsETo_def) apply (blast intro: leadsETo_weaken_L) done lemma LeadsETo_weaken: - "[| F : A LeadsTo[CC'] A'; + "[| F \ A LeadsTo[CC'] A'; B <= A; A' <= B'; CC' <= CC |] - ==> F : B LeadsTo[CC] B'" + ==> F \ B LeadsTo[CC] B'" apply (simp (no_asm_use) add: LeadsETo_def) apply (blast intro: leadsETo_weaken) done @@ -445,12 +445,12 @@ (*Postcondition can be strengthened to (reachable F Int B) *) lemma reachable_ensures: - "F : A ensures B ==> F : (reachable F Int A) ensures B" + "F \ A ensures B \ F \ (reachable F Int A) ensures B" apply (rule stable_ensures_Int [THEN ensures_weaken_R], auto) done lemma lel_lemma: - "F : A leadsTo B ==> F : (reachable F Int A) leadsTo[Pow(reachable F)] B" + "F \ A leadsTo B \ F \ (reachable F Int A) leadsTo[Pow(reachable F)] B" apply (erule leadsTo_induct) apply (blast intro: reachable_ensures) apply (blast dest: e_psp_stable2 intro: leadsETo_Trans leadsETo_weaken_L) @@ -483,13 +483,13 @@ by (simp add: givenBy_eq_Collect, best) lemma extend_set_givenBy_I: - "D : givenBy v ==> extend_set h D : givenBy (v o f)" + "D \ givenBy v ==> extend_set h D \ givenBy (v o f)" apply (simp (no_asm_use) add: givenBy_eq_all, blast) done lemma leadsETo_imp_extend_leadsETo: - "F : A leadsTo[CC] B - ==> extend h F : (extend_set h A) leadsTo[extend_set h ` CC] + "F \ A leadsTo[CC] B + ==> extend h F \ (extend_set h A) leadsTo[extend_set h ` CC] (extend_set h B)" apply (erule leadsETo_induct) apply (force intro: subset_imp_ensures @@ -502,11 +502,11 @@ (*This version's stronger in the "ensures" precondition BUT there's no ensures_weaken_L*) lemma Join_project_ensures_strong: - "[| project h C G ~: transient (project_set h C Int (A-B)) | + "[| project h C G \ transient (project_set h C Int (A-B)) | project_set h C Int (A - B) = {}; - extend h F\G : stable C; - F\project h C G : (project_set h C Int A) ensures B |] - ==> extend h F\G : (C Int extend_set h A) ensures (extend_set h B)" + extend h F\G \ stable C; + F\project h C G \ (project_set h C Int A) ensures B |] + ==> extend h F\G \ (C Int extend_set h A) ensures (extend_set h B)" apply (subst Int_extend_set_lemma [symmetric]) apply (rule Join_project_ensures) apply (auto simp add: Int_Diff) @@ -595,22 +595,22 @@ (*Lemma for the Trans case*) lemma pli_lemma: - "[| extend h F\G : stable C; + "[| extend h F\G \ stable C; F\project h C G - : project_set h C Int project_set h A leadsTo project_set h B |] + \ project_set h C Int project_set h A leadsTo project_set h B |] ==> F\project h C G - : project_set h C Int project_set h A leadsTo + \ project_set h C Int project_set h A leadsTo project_set h C Int project_set h B" apply (rule psp_stable2 [THEN leadsTo_weaken_L]) apply (auto simp add: project_stable_project_set extend_stable_project_set) done lemma project_leadsETo_I_lemma: - "[| extend h F\G : stable C; - extend h F\G : + "[| extend h F\G \ stable C; + extend h F\G \ (C Int A) leadsTo[(%D. C Int D)`givenBy f] B |] ==> F\project h C G - : (project_set h C Int project_set h (C Int A)) leadsTo (project_set h B)" + \ (project_set h C Int project_set h (C Int A)) leadsTo (project_set h B)" apply (erule leadsETo_induct) prefer 3 apply (simp only: Int_UN_distrib project_set_Union) @@ -622,22 +622,22 @@ done lemma project_leadsETo_I: - "extend h F\G : (extend_set h A) leadsTo[givenBy f] (extend_set h B) - ==> F\project h UNIV G : A leadsTo B" + "extend h F\G \ (extend_set h A) leadsTo[givenBy f] (extend_set h B) + \ F\project h UNIV G \ A leadsTo B" apply (rule project_leadsETo_I_lemma [THEN leadsTo_weaken], auto) done lemma project_LeadsETo_I: - "extend h F\G : (extend_set h A) LeadsTo[givenBy f] (extend_set h B) - ==> F\project h (reachable (extend h F\G)) G - : A LeadsTo B" + "extend h F\G \ (extend_set h A) LeadsTo[givenBy f] (extend_set h B) + \ F\project h (reachable (extend h F\G)) G + \ A LeadsTo B" apply (simp (no_asm_use) add: LeadsTo_def LeadsETo_def) apply (rule project_leadsETo_I_lemma [THEN leadsTo_weaken]) apply (auto simp add: project_set_reachable_extend_eq [symmetric]) done lemma projecting_leadsTo: - "projecting (%G. UNIV) h F + "projecting (\G. UNIV) h F (extend_set h A leadsTo[givenBy f] extend_set h B) (A leadsTo B)" apply (unfold projecting_def) @@ -645,7 +645,7 @@ done lemma projecting_LeadsTo: - "projecting (%G. reachable (extend h F\G)) h F + "projecting (\G. reachable (extend h F\G)) h F (extend_set h A LeadsTo[givenBy f] extend_set h B) (A LeadsTo B)" apply (unfold projecting_def) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Extend.thy --- a/src/HOL/UNITY/Extend.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Extend.thy Thu Feb 15 13:04:36 2018 +0100 @@ -49,7 +49,7 @@ where "project h C F = mk_program (project_set h (Init F), project_act h ` Restrict C ` Acts F, - {act. Restrict (project_set h C) act : + {act. Restrict (project_set h C) act \ project_act h ` Restrict C ` AllowedActs F})" locale Extend = @@ -70,7 +70,7 @@ subsection\Restrict\ (*MOVE to Relation.thy?*) -lemma Restrict_iff [iff]: "((x,y): Restrict A r) = ((x,y): r & x \ A)" +lemma Restrict_iff [iff]: "((x,y) \ Restrict A r) = ((x,y) \ r & x \ A)" by (unfold Restrict_def, blast) lemma Restrict_UNIV [simp]: "Restrict UNIV = id" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/FP.thy --- a/src/HOL/UNITY/FP.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/FP.thy Thu Feb 15 13:04:36 2018 +0100 @@ -10,21 +10,21 @@ theory FP imports UNITY begin definition FP_Orig :: "'a program => 'a set" where - "FP_Orig F == \{A. ALL B. F : stable (A Int B)}" + "FP_Orig F == \{A. \B. F \ stable (A \ B)}" definition FP :: "'a program => 'a set" where - "FP F == {s. F : stable {s}}" + "FP F == {s. F \ stable {s}}" -lemma stable_FP_Orig_Int: "F : stable (FP_Orig F Int B)" +lemma stable_FP_Orig_Int: "F \ stable (FP_Orig F Int B)" apply (simp only: FP_Orig_def stable_def Int_Union2) apply (blast intro: constrains_UN) done lemma FP_Orig_weakest: - "(!!B. F : stable (A Int B)) ==> A <= FP_Orig F" + "(\B. F \ stable (A \ B)) \ A <= FP_Orig F" by (simp add: FP_Orig_def stable_def, blast) -lemma stable_FP_Int: "F : stable (FP F Int B)" +lemma stable_FP_Int: "F \ stable (FP F \ B)" apply (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x}) ") prefer 2 apply blast apply (simp (no_asm_simp) add: Int_insert_right) @@ -42,7 +42,7 @@ done lemma FP_weakest: - "(!!B. F : stable (A Int B)) ==> A <= FP F" + "(\B. F \ stable (A Int B)) \ A <= FP F" by (simp add: FP_equivalence FP_Orig_weakest) lemma Compl_FP: diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Guar.thy --- a/src/HOL/UNITY/Guar.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Guar.thy Thu Feb 15 13:04:36 2018 +0100 @@ -85,8 +85,8 @@ done lemma ex2: - "\GG. finite GG & GG \ X \ {} --> OK GG (%G. G) -->(\G \ GG. G):X - ==> ex_prop X" + "\GG. finite GG & GG \ X \ {} \ OK GG (\G. G) \ (\G \ GG. G) \ X + \ ex_prop X" apply (unfold ex_prop_def, clarify) apply (drule_tac x = "{F,G}" in spec) apply (auto dest: ok_sym simp add: OK_iff_ok) @@ -140,7 +140,7 @@ (*Chandy & Sanders take this as a definition*) lemma uv_prop_finite: "uv_prop X = - (\GG. finite GG & GG \ X & OK GG (%G. G) --> (\G \ GG. G): X)" + (\GG. finite GG \ GG \ X \ OK GG (\G. G) \ (\G \ GG. G) \ X)" by (blast intro: uv1 uv2) subsection\Guarantees\ @@ -403,8 +403,8 @@ by (simp add: wg_equiv) lemma wg_finite: - "\FF. finite FF & FF \ X \ {} --> OK FF (%F. F) - --> (\F\FF. ((\F \ FF. F): wg F X) = ((\F \ FF. F):X))" + "\FF. finite FF \ FF \ X \ {} \ OK FF (\F. F) + \ (\F\FF. ((\F \ FF. F) \ wg F X) = ((\F \ FF. F) \ X))" apply clarify apply (subgoal_tac "F component_of (\F \ FF. F) ") apply (drule_tac X = X in component_of_wg, simp) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/ListOrder.thy --- a/src/HOL/UNITY/ListOrder.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/ListOrder.thy Thu Feb 15 13:04:36 2018 +0100 @@ -20,18 +20,18 @@ genPrefix :: "('a * 'a)set => ('a list * 'a list)set" for r :: "('a * 'a)set" where - Nil: "([],[]) : genPrefix(r)" + Nil: "([],[]) \ genPrefix(r)" - | prepend: "[| (xs,ys) : genPrefix(r); (x,y) : r |] ==> - (x#xs, y#ys) : genPrefix(r)" + | prepend: "[| (xs,ys) \ genPrefix(r); (x,y) \ r |] ==> + (x#xs, y#ys) \ genPrefix(r)" - | append: "(xs,ys) : genPrefix(r) ==> (xs, ys@zs) : genPrefix(r)" + | append: "(xs,ys) \ genPrefix(r) ==> (xs, ys@zs) \ genPrefix(r)" instantiation list :: (type) ord begin definition - prefix_def: "xs <= zs \ (xs, zs) : genPrefix Id" + prefix_def: "xs <= zs \ (xs, zs) \ genPrefix Id" definition strict_prefix_def: "xs < zs \ xs \ zs \ \ zs \ (xs :: 'a list)" @@ -50,37 +50,37 @@ abbreviation pfixLe :: "[nat list, nat list] => bool" (infixl "pfixLe" 50) where - "xs pfixLe ys == (xs,ys) : genPrefix Le" + "xs pfixLe ys == (xs,ys) \ genPrefix Le" abbreviation pfixGe :: "[nat list, nat list] => bool" (infixl "pfixGe" 50) where - "xs pfixGe ys == (xs,ys) : genPrefix Ge" + "xs pfixGe ys == (xs,ys) \ genPrefix Ge" subsection\preliminary lemmas\ -lemma Nil_genPrefix [iff]: "([], xs) : genPrefix r" +lemma Nil_genPrefix [iff]: "([], xs) \ genPrefix r" by (cut_tac genPrefix.Nil [THEN genPrefix.append], auto) -lemma genPrefix_length_le: "(xs,ys) : genPrefix r ==> length xs <= length ys" +lemma genPrefix_length_le: "(xs,ys) \ genPrefix r \ length xs <= length ys" by (erule genPrefix.induct, auto) lemma cdlemma: - "[| (xs', ys'): genPrefix r |] - ==> (ALL x xs. xs' = x#xs --> (EX y ys. ys' = y#ys & (x,y) : r & (xs, ys) : genPrefix r))" + "[| (xs', ys') \ genPrefix r |] + ==> (\x xs. xs' = x#xs \ (\y ys. ys' = y#ys & (x,y) \ r & (xs, ys) \ genPrefix r))" apply (erule genPrefix.induct, blast, blast) apply (force intro: genPrefix.append) done (*As usual converting it to an elimination rule is tiresome*) lemma cons_genPrefixE [elim!]: - "[| (x#xs, zs): genPrefix r; - !!y ys. [| zs = y#ys; (x,y) : r; (xs, ys) : genPrefix r |] ==> P + "[| (x#xs, zs) \ genPrefix r; + !!y ys. [| zs = y#ys; (x,y) \ r; (xs, ys) \ genPrefix r |] ==> P |] ==> P" by (drule cdlemma, simp, blast) lemma Cons_genPrefix_Cons [iff]: - "((x#xs,y#ys) : genPrefix r) = ((x,y) : r & (xs,ys) : genPrefix r)" + "((x#xs,y#ys) \ genPrefix r) = ((x,y) \ r \ (xs,ys) \ genPrefix r)" by (blast intro: genPrefix.prepend) @@ -93,7 +93,7 @@ apply (blast intro: genPrefix.Nil) done -lemma genPrefix_refl [simp]: "refl r ==> (l,l) : genPrefix r" +lemma genPrefix_refl [simp]: "refl r \ (l,l) \ genPrefix r" by (erule refl_onD [OF refl_genPrefix UNIV_I]) lemma genPrefix_mono: "r<=s ==> genPrefix r <= genPrefix s" @@ -107,13 +107,13 @@ (*A lemma for proving genPrefix_trans_O*) lemma append_genPrefix: - "(xs @ ys, zs) : genPrefix r \ (xs, zs) : genPrefix r" + "(xs @ ys, zs) \ genPrefix r \ (xs, zs) \ genPrefix r" by (induct xs arbitrary: zs) auto (*Lemma proving transitivity and more*) lemma genPrefix_trans_O: - assumes "(x, y) : genPrefix r" - shows "\z. (y, z) : genPrefix s \ (x, z) : genPrefix (r O s)" + assumes "(x, y) \ genPrefix r" + shows "\z. (y, z) \ genPrefix s \ (x, z) \ genPrefix (r O s)" apply (atomize (full)) using assms apply induct @@ -123,22 +123,22 @@ done lemma genPrefix_trans: - "(x, y) : genPrefix r \ (y, z) : genPrefix r \ trans r - \ (x, z) : genPrefix r" + "(x, y) \ genPrefix r \ (y, z) \ genPrefix r \ trans r + \ (x, z) \ genPrefix r" apply (rule trans_O_subset [THEN genPrefix_mono, THEN subsetD]) apply assumption apply (blast intro: genPrefix_trans_O) done lemma prefix_genPrefix_trans: - "[| x<=y; (y,z) : genPrefix r |] ==> (x, z) : genPrefix r" + "[| x<=y; (y,z) \ genPrefix r |] ==> (x, z) \ genPrefix r" apply (unfold prefix_def) apply (drule genPrefix_trans_O, assumption) apply simp done lemma genPrefix_prefix_trans: - "[| (x,y) : genPrefix r; y<=z |] ==> (x,z) : genPrefix r" + "[| (x,y) \ genPrefix r; y<=z |] ==> (x,z) \ genPrefix r" apply (unfold prefix_def) apply (drule genPrefix_trans_O, assumption) apply simp @@ -151,9 +151,9 @@ (** Antisymmetry **) lemma genPrefix_antisym: - assumes 1: "(xs, ys) : genPrefix r" + assumes 1: "(xs, ys) \ genPrefix r" and 2: "antisym r" - and 3: "(ys, xs) : genPrefix r" + and 3: "(ys, xs) \ genPrefix r" shows "xs = ys" using 1 3 proof induct @@ -178,29 +178,29 @@ subsection\recursion equations\ -lemma genPrefix_Nil [simp]: "((xs, []) : genPrefix r) = (xs = [])" +lemma genPrefix_Nil [simp]: "((xs, []) \ genPrefix r) = (xs = [])" by (induct xs) auto lemma same_genPrefix_genPrefix [simp]: - "refl r ==> ((xs@ys, xs@zs) : genPrefix r) = ((ys,zs) : genPrefix r)" + "refl r \ ((xs@ys, xs@zs) \ genPrefix r) = ((ys,zs) \ genPrefix r)" by (induct xs) (simp_all add: refl_on_def) lemma genPrefix_Cons: - "((xs, y#ys) : genPrefix r) = - (xs=[] | (EX z zs. xs=z#zs & (z,y) : r & (zs,ys) : genPrefix r))" + "((xs, y#ys) \ genPrefix r) = + (xs=[] | (\z zs. xs=z#zs & (z,y) \ r & (zs,ys) \ genPrefix r))" by (cases xs) auto lemma genPrefix_take_append: - "[| refl r; (xs,ys) : genPrefix r |] - ==> (xs@zs, take (length xs) ys @ zs) : genPrefix r" + "[| refl r; (xs,ys) \ genPrefix r |] + ==> (xs@zs, take (length xs) ys @ zs) \ genPrefix r" apply (erule genPrefix.induct) apply (frule_tac [3] genPrefix_length_le) apply (simp_all (no_asm_simp) add: diff_is_0_eq [THEN iffD2]) done lemma genPrefix_append_both: - "[| refl r; (xs,ys) : genPrefix r; length xs = length ys |] - ==> (xs@zs, ys @ zs) : genPrefix r" + "[| refl r; (xs,ys) \ genPrefix r; length xs = length ys |] + ==> (xs@zs, ys @ zs) \ genPrefix r" apply (drule genPrefix_take_append, assumption) apply simp done @@ -211,8 +211,8 @@ by auto lemma aolemma: - "[| (xs,ys) : genPrefix r; refl r |] - ==> length xs < length ys --> (xs @ [ys ! length xs], ys) : genPrefix r" + "[| (xs,ys) \ genPrefix r; refl r |] + ==> length xs < length ys \ (xs @ [ys ! length xs], ys) \ genPrefix r" apply (erule genPrefix.induct) apply blast apply simp @@ -226,15 +226,15 @@ done lemma append_one_genPrefix: - "[| (xs,ys) : genPrefix r; length xs < length ys; refl r |] - ==> (xs @ [ys ! length xs], ys) : genPrefix r" + "[| (xs,ys) \ genPrefix r; length xs < length ys; refl r |] + ==> (xs @ [ys ! length xs], ys) \ genPrefix r" by (blast intro: aolemma [THEN mp]) (** Proving the equivalence with Charpentier's definition **) lemma genPrefix_imp_nth: - "i < length xs \ (xs, ys) : genPrefix r \ (xs ! i, ys ! i) : r" + "i < length xs \ (xs, ys) \ genPrefix r \ (xs ! i, ys ! i) \ r" apply (induct xs arbitrary: i ys) apply auto apply (case_tac i) @@ -243,8 +243,8 @@ lemma nth_imp_genPrefix: "length xs <= length ys \ - (\i. i < length xs --> (xs ! i, ys ! i) : r) \ - (xs, ys) : genPrefix r" + (\i. i < length xs \ (xs ! i, ys ! i) \ r) \ + (xs, ys) \ genPrefix r" apply (induct xs arbitrary: ys) apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib) apply (case_tac ys) @@ -252,8 +252,8 @@ done lemma genPrefix_iff_nth: - "((xs,ys) : genPrefix r) = - (length xs <= length ys & (ALL i. i < length xs --> (xs!i, ys!i) : r))" + "((xs,ys) \ genPrefix r) = + (length xs <= length ys & (\i. i < length xs \ (xs!i, ys!i) \ r))" apply (blast intro: genPrefix_length_le genPrefix_imp_nth nth_imp_genPrefix) done @@ -315,7 +315,7 @@ done lemma prefix_Cons: - "(xs <= y#ys) = (xs=[] | (? zs. xs=y#zs & zs <= ys))" + "(xs <= y#ys) = (xs=[] | (\zs. xs=y#zs \ zs <= ys))" by (simp add: prefix_def genPrefix_Cons) lemma append_one_prefix: @@ -343,7 +343,7 @@ by (blast intro: monoI prefix_length_le) (*Equivalence to the definition used in Lex/Prefix.thy*) -lemma prefix_iff: "(xs <= zs) = (EX ys. zs = xs@ys)" +lemma prefix_iff: "(xs <= zs) = (\ys. zs = xs@ys)" apply (unfold prefix_def) apply (auto simp add: genPrefix_iff_nth nth_append) apply (rule_tac x = "drop (length xs) zs" in exI) @@ -363,7 +363,7 @@ done lemma prefix_append_iff: - "(xs <= ys@zs) = (xs <= ys | (? us. xs = ys@us & us <= zs))" + "(xs <= ys@zs) = (xs <= ys | (\us. xs = ys@us & us <= zs))" apply (rule_tac xs = zs in rev_induct) apply force apply (simp del: append_assoc add: append_assoc [symmetric], force) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/PPROD.thy --- a/src/HOL/UNITY/PPROD.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/PPROD.thy Thu Feb 15 13:04:36 2018 +0100 @@ -120,8 +120,8 @@ something like \lift_preserves_sub\ to rewrite the third. However there's no obvious alternative for the third premise.\ lemma guarantees_PLam_I: - "[| lift i (F i): X guarantees Y; i \ I; - OK I (%i. lift i (F i)) |] + "[| lift i (F i) \ X guarantees Y; i \ I; + OK I (\i. lift i (F i)) |] ==> (PLam I F) \ X guarantees Y" apply (unfold PLam_def) apply (simp add: guarantees_JN_I) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Simple/NSP_Bad.thy --- a/src/HOL/UNITY/Simple/NSP_Bad.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Thu Feb 15 13:04:36 2018 +0100 @@ -258,7 +258,7 @@ "[| A \ bad; B \ bad |] ==> Nprg \ Always {s. Says B A (Crypt (pubK A) \Nonce NA, Nonce NB\) \ set s & - (ALL C. Says A C (Crypt (pubK C) (Nonce NB)) \ set s) + (\C. Says A C (Crypt (pubK C) (Nonce NB)) \ set s) --> Nonce NB \ analz (spies s)}" apply ns_induct apply (simp_all (no_asm_simp) add: all_conj_distrib) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Simple/Reach.thy --- a/src/HOL/UNITY/Simple/Reach.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Simple/Reach.thy Thu Feb 15 13:04:36 2018 +0100 @@ -24,7 +24,7 @@ where "Rprg = mk_total_program ({%v. v=init}, \(u,v)\edges. {asgt u v}, UNIV)" definition reach_invariant :: "state set" - where "reach_invariant = {s. (\v. s v --> (init, v) \ edges^*) & s init}" + where "reach_invariant = {s. (\v. s v --> (init, v) \ edges\<^sup>*) & s init}" definition fixedpoint :: "state set" where "fixedpoint = {s. \(u,v)\edges. s u --> s v}" @@ -68,7 +68,7 @@ (*If it reaches a fixedpoint, it has found a solution*) lemma fixedpoint_invariant_correct: - "fixedpoint \ reach_invariant = { %v. (init, v) \ edges^* }" + "fixedpoint \ reach_invariant = { %v. (init, v) \ edges\<^sup>* }" apply (unfold fixedpoint_def) apply (rule equalityI) apply (auto intro!: ext) @@ -152,7 +152,7 @@ apply (rule LeadsTo_Un_fixedpoint) done -lemma LeadsTo_correct: "Rprg \ UNIV LeadsTo { %v. (init, v) \ edges^* }" +lemma LeadsTo_correct: "Rprg \ UNIV LeadsTo { %v. (init, v) \ edges\<^sup>* }" apply (subst fixedpoint_invariant_correct [symmetric]) apply (rule Always_LeadsTo_weaken [OF reach_invariant LeadsTo_fixedpoint], auto) done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Simple/Reachability.thy --- a/src/HOL/UNITY/Simple/Reachability.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Simple/Reachability.thy Thu Feb 15 13:04:36 2018 +0100 @@ -296,7 +296,7 @@ lemma Leadsto_reachability_AND_nmsg_0: "[| v \ V; w \ V |] ==> F \ UNIV LeadsTo - ((reachable v <==> {s. (root,v): REACHABLE}) \ nmsg_eq 0 (v,w))" + ((reachable v <==> {s. (root,v) \ REACHABLE}) \ nmsg_eq 0 (v,w))" apply (rule LeadsTo_Reachability [THEN LeadsTo_Trans], blast) apply (subgoal_tac "F \ (reachable v <==> {s. (root,v) \ REACHABLE}) \ @@ -355,7 +355,7 @@ apply simp apply (subgoal_tac "({s. (s \ reachable v) = ((root,v) \ REACHABLE) }) = - ({s. (s \ reachable v) = ( (root,v) \ REACHABLE) } Int + ({s. (s \ reachable v) = ( (root,v) \ REACHABLE) } \ (- {} \ nmsg_eq 0 (v,w)))") apply blast+ done diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/SubstAx.thy --- a/src/HOL/UNITY/SubstAx.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/SubstAx.thy Thu Feb 15 13:04:36 2018 +0100 @@ -304,7 +304,7 @@ lemma LeadsTo_wf_induct: "[| wf r; \m. F \ (A \ f-`{m}) LeadsTo - ((A \ f-`(r^-1 `` {m})) \ B) |] + ((A \ f-`(r\ `` {m})) \ B) |] ==> F \ A LeadsTo B" apply (simp add: LeadsTo_eq_leadsTo) apply (erule leadsTo_wf_induct) @@ -315,7 +315,7 @@ lemma Bounded_induct: "[| wf r; \m \ I. F \ (A \ f-`{m}) LeadsTo - ((A \ f-`(r^-1 `` {m})) \ B) |] + ((A \ f-`(r\ `` {m})) \ B) |] ==> F \ A LeadsTo ((A - (f-`I)) \ B)" apply (erule LeadsTo_wf_induct, safe) apply (case_tac "m \ I") diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Transformers.thy --- a/src/HOL/UNITY/Transformers.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Transformers.thy Thu Feb 15 13:04:36 2018 +0100 @@ -22,7 +22,7 @@ definition wp :: "[('a*'a) set, 'a set] => 'a set" where \ \Dijkstra's weakest-precondition operator (for an individual command)\ - "wp act B == - (act^-1 `` (-B))" + "wp act B == - (act\ `` (-B))" definition awp :: "['a program, 'a set] => 'a set" where \ \Dijkstra's weakest-precondition operator (for a program)\ diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/UNITY.thy --- a/src/HOL/UNITY/UNITY.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/UNITY.thy Thu Feb 15 13:04:36 2018 +0100 @@ -15,7 +15,7 @@ definition "Program = {(init:: 'a set, acts :: ('a * 'a)set set, - allowed :: ('a * 'a)set set). Id \ acts & Id: allowed}" + allowed :: ('a * 'a)set set). Id \ acts & Id \ allowed}" typedef 'a program = "Program :: ('a set * ('a * 'a) set set * ('a * 'a) set set) set" morphisms Rep_Program Abs_Program @@ -127,12 +127,12 @@ subsubsection\co\ lemma constrainsI: - "(!!act s s'. [| act: Acts F; (s,s') \ act; s \ A |] ==> s': A') + "(!!act s s'. [| act \ Acts F; (s,s') \ act; s \ A |] ==> s' \ A') ==> F \ A co A'" by (simp add: constrains_def, blast) lemma constrainsD: - "[| F \ A co A'; act: Acts F; (s,s'): act; s \ A |] ==> s': A'" + "[| F \ A co A'; act \ Acts F; (s,s') \ act; s \ A |] ==> s' \ A'" by (unfold constrains_def, blast) lemma constrains_empty [iff]: "F \ {} co B" @@ -351,7 +351,7 @@ lemma Image_less_than [simp]: "less_than `` {k} = greaterThan k" by blast -lemma Image_inverse_less_than [simp]: "less_than^-1 `` {k} = lessThan k" +lemma Image_inverse_less_than [simp]: "less_than\ `` {k} = lessThan k" by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/Union.thy --- a/src/HOL/UNITY/Union.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/Union.thy Thu Feb 15 13:04:36 2018 +0100 @@ -36,7 +36,7 @@ (*Characterizes safety properties. Used with specifying Allowed*) definition safety_prop :: "'a program set => bool" - where "safety_prop X \ SKIP: X & (\G. Acts G \ UNION X Acts --> G \ X)" + where "safety_prop X \ SKIP \ X \ (\G. Acts G \ UNION X Acts \ G \ X)" syntax "_JOIN1" :: "[pttrns, 'b set] => 'b set" ("(3\_./ _)" 10) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/UNITY/WFair.thy --- a/src/HOL/UNITY/WFair.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/UNITY/WFair.thy Thu Feb 15 13:04:36 2018 +0100 @@ -91,12 +91,12 @@ done lemma transientI: - "[| act: Acts F; act``A \ -A |] ==> F \ transient A" + "[| act \ Acts F; act``A \ -A |] ==> F \ transient A" by (unfold transient_def, blast) lemma transientE: "[| F \ transient A; - !!act. [| act: Acts F; act``A \ -A |] ==> P |] + \act. [| act \ Acts F; act``A \ -A |] ==> P |] ==> P" by (unfold transient_def, blast) @@ -115,7 +115,7 @@ done lemma totalize_transientI: - "[| act: Acts F; A \ Domain act; act``A \ -A |] + "[| act \ Acts F; A \ Domain act; act``A \ -A |] ==> totalize F \ transient A" by (simp add: totalize_transient_iff, blast) @@ -420,10 +420,10 @@ lemma leadsTo_wf_induct_lemma: "[| wf r; \m. F \ (A \ f-`{m}) leadsTo - ((A \ f-`(r^-1 `` {m})) \ B) |] + ((A \ f-`(r\ `` {m})) \ B) |] ==> F \ (A \ f-`{m}) leadsTo B" apply (erule_tac a = m in wf_induct) -apply (subgoal_tac "F \ (A \ (f -` (r^-1 `` {x}))) leadsTo B") +apply (subgoal_tac "F \ (A \ (f -` (r\ `` {x}))) leadsTo B") apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate) apply (subst vimage_eq_UN) apply (simp only: UN_simps [symmetric]) @@ -435,7 +435,7 @@ lemma leadsTo_wf_induct: "[| wf r; \m. F \ (A \ f-`{m}) leadsTo - ((A \ f-`(r^-1 `` {m})) \ B) |] + ((A \ f-`(r\ `` {m})) \ B) |] ==> F \ A leadsTo B" apply (rule_tac t = A in subst) defer 1 @@ -449,7 +449,7 @@ lemma bounded_induct: "[| wf r; \m \ I. F \ (A \ f-`{m}) leadsTo - ((A \ f-`(r^-1 `` {m})) \ B) |] + ((A \ f-`(r\ `` {m})) \ B) |] ==> F \ A leadsTo ((A - (f-`I)) \ B)" apply (erule leadsTo_wf_induct, safe) apply (case_tac "m \ I") diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Word/Misc_Typedef.thy --- a/src/HOL/Word/Misc_Typedef.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Word/Misc_Typedef.thy Thu Feb 15 13:04:36 2018 +0100 @@ -124,7 +124,7 @@ lemmas td = td_thm -lemma set_iff_norm: "w : A \ w = norm w" +lemma set_iff_norm: "w \ A \ w = norm w" by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) lemma inverse_norm: "Abs n = w \ Rep w = norm n" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ZF/Games.thy --- a/src/HOL/ZF/Games.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ZF/Games.thy Thu Feb 15 13:04:36 2018 +0100 @@ -85,18 +85,18 @@ apply (rule_tac x=r in exI, insert Elem_Opair_exists, blast) done -lemma is_option_of_subset_is_Elem_of: "is_option_of \ (is_Elem_of^+)" +lemma is_option_of_subset_is_Elem_of: "is_option_of \ (is_Elem_of\<^sup>+)" proof - { fix opt fix g assume "(opt, g) \ is_option_of" - then have "\ u v. (opt, u) \ (is_Elem_of^+) \ (u,v) \ (is_Elem_of^+) \ (v,g) \ (is_Elem_of^+)" + then have "\ u v. (opt, u) \ (is_Elem_of\<^sup>+) \ (u,v) \ (is_Elem_of\<^sup>+) \ (v,g) \ (is_Elem_of\<^sup>+)" apply - apply (drule option2elem) apply (auto simp add: r_into_trancl' is_Elem_of_def) done - then have "(opt, g) \ (is_Elem_of^+)" + then have "(opt, g) \ (is_Elem_of\<^sup>+)" by (blast intro: trancl_into_rtrancl trancl_rtrancl_trancl) } then show ?thesis by auto @@ -104,7 +104,7 @@ lemma wfzf_is_option_of: "wfzf is_option_of" proof - - have "wfzf (is_Elem_of^+)" by (simp add: wfzf_trancl wfzf_is_Elem_of) + have "wfzf (is_Elem_of\<^sup>+)" by (simp add: wfzf_trancl wfzf_is_Elem_of) then show ?thesis apply (rule wfzf_subset) apply (rule is_option_of_subset_is_Elem_of) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ZF/HOLZF.thy --- a/src/HOL/ZF/HOLZF.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ZF/HOLZF.thy Thu Feb 15 13:04:36 2018 +0100 @@ -32,25 +32,25 @@ "SucNat x == union x (Singleton x)" definition subset :: "ZF \ ZF \ bool" where - "subset A B == ! x. Elem x A \ Elem x B" + "subset A B \ \x. Elem x A \ Elem x B" axiomatization where Empty: "Not (Elem x Empty)" and - Ext: "(x = y) = (! z. Elem z x = Elem z y)" and - Sum: "Elem z (Sum x) = (? y. Elem z y & Elem y x)" and + Ext: "(x = y) = (\z. Elem z x = Elem z y)" and + Sum: "Elem z (Sum x) = (\y. Elem z y \ Elem y x)" and Power: "Elem y (Power x) = (subset y x)" and - Repl: "Elem b (Repl A f) = (? a. Elem a A & b = f a)" and - Regularity: "A \ Empty \ (? x. Elem x A & (! y. Elem y x \ Not (Elem y A)))" and - Infinity: "Elem Empty Inf & (! x. Elem x Inf \ Elem (SucNat x) Inf)" + Repl: "Elem b (Repl A f) = (\a. Elem a A \ b = f a)" and + Regularity: "A \ Empty \ (\x. Elem x A \ (\y. Elem y x \ Not (Elem y A)))" and + Infinity: "Elem Empty Inf \ (\x. Elem x Inf \ Elem (SucNat x) Inf)" definition Sep :: "ZF \ (ZF \ bool) \ ZF" where - "Sep A p == (if (!x. Elem x A \ Not (p x)) then Empty else + "Sep A p == (if (\x. Elem x A \ Not (p x)) then Empty else (let z = (\ x. Elem x A & p x) in - let f = % x. (if p x then x else z) in Repl A f))" + let f = \x. (if p x then x else z) in Repl A f))" thm Power[unfolded subset_def] -theorem Sep: "Elem b (Sep A p) = (Elem b A & p b)" +theorem Sep: "Elem b (Sep A p) = (Elem b A \ p b)" apply (auto simp add: Sep_def Empty) apply (auto simp add: Let_def Repl) apply (rule someI2, auto)+ @@ -59,7 +59,7 @@ lemma subset_empty: "subset Empty A" by (simp add: subset_def Empty) -theorem Upair: "Elem x (Upair a b) = (x = a | x = b)" +theorem Upair: "Elem x (Upair a b) = (x = a \ x = b)" apply (auto simp add: Upair_def Repl) apply (rule exI[where x=Empty]) apply (simp add: Power subset_empty) @@ -103,14 +103,14 @@ definition Replacement :: "ZF \ (ZF \ ZF option) \ ZF" where "Replacement A f == Repl (Sep A (% a. f a \ None)) (the o f)" -theorem Replacement: "Elem y (Replacement A f) = (? x. Elem x A & f x = Some y)" +theorem Replacement: "Elem y (Replacement A f) = (\x. Elem x A \ f x = Some y)" by (auto simp add: Replacement_def Repl Sep) definition Fst :: "ZF \ ZF" where - "Fst q == SOME x. ? y. q = Opair x y" + "Fst q == SOME x. \y. q = Opair x y" definition Snd :: "ZF \ ZF" where - "Snd q == SOME y. ? x. q = Opair x y" + "Snd q == SOME y. \x. q = Opair x y" theorem Fst: "Fst (Opair x y) = x" apply (simp add: Fst_def) @@ -125,7 +125,7 @@ done definition isOpair :: "ZF \ bool" where - "isOpair q == ? x y. q = Opair x y" + "isOpair q == \x y. q = Opair x y" lemma isOpair: "isOpair (Opair x y) = True" by (auto simp add: isOpair_def) @@ -136,7 +136,7 @@ definition CartProd :: "ZF \ ZF \ ZF" where "CartProd A B == Sum(Repl A (% a. Repl B (% b. Opair a b)))" -lemma CartProd: "Elem x (CartProd A B) = (? a b. Elem a A & Elem b B & x = (Opair a b))" +lemma CartProd: "Elem x (CartProd A B) = (\a b. Elem a A \ Elem b B \ x = (Opair a b))" apply (auto simp add: CartProd_def Sum Repl) apply (rule_tac x="Repl B (Opair a)" in exI) apply (auto simp add: Repl) @@ -166,7 +166,7 @@ definition Range :: "ZF \ ZF" where "Range f == Replacement f (% p. if isOpair p then Some (Snd p) else None)" -theorem Domain: "Elem x (Domain f) = (? y. Elem (Opair x y) f)" +theorem Domain: "Elem x (Domain f) = (\y. Elem (Opair x y) f)" apply (auto simp add: Domain_def Replacement) apply (rule_tac x="Snd xa" in exI) apply (simp add: FstSnd) @@ -174,7 +174,7 @@ apply (simp add: isOpair Fst) done -theorem Range: "Elem y (Range f) = (? x. Elem (Opair x y) f)" +theorem Range: "Elem y (Range f) = (\x. Elem (Opair x y) f)" apply (auto simp add: Range_def Replacement) apply (rule_tac x="Fst x" in exI) apply (simp add: FstSnd) @@ -192,7 +192,7 @@ "f \ x == (THE y. Elem (Opair x y) f)" definition isFun :: "ZF \ bool" where - "isFun f == (! x y1 y2. Elem (Opair x y1) f & Elem (Opair x y2) f \ y1 = y2)" + "isFun f == (\x y1 y2. Elem (Opair x y1) f & Elem (Opair x y2) f \ y1 = y2)" definition Lambda :: "ZF \ (ZF \ ZF) \ ZF" where "Lambda A f == Repl A (% x. Opair x (f x))" @@ -214,7 +214,7 @@ apply (auto simp add: Fst isOpair_def) done -lemma Lambda_ext: "(Lambda s f = Lambda t g) = (s = t & (! x. Elem x s \ f x = g x))" +lemma Lambda_ext: "(Lambda s f = Lambda t g) = (s = t \ (\x. Elem x s \ f x = g x))" proof - have "Lambda s f = Lambda t g \ s = t" apply (subst domain_Lambda[where A = s and f = f, symmetric]) @@ -282,13 +282,13 @@ apply (auto simp add: the_equality) done -lemma fun_range_witness: "\isFun f; Elem y (Range f)\ \ ? x. Elem x (Domain f) & f\x = y" +lemma fun_range_witness: "\isFun f; Elem y (Range f)\ \ \x. Elem x (Domain f) & f\x = y" apply (auto simp add: Range) apply (rule_tac x="x" in exI) apply (auto simp add: app_def the_equality isFun_def Domain) done -lemma Elem_Fun_Lambda: "Elem F (Fun U V) \ ? f. F = Lambda U f" +lemma Elem_Fun_Lambda: "Elem F (Fun U V) \ \f. F = Lambda U f" apply (rule exI[where x= "% x. (THE y. Elem (Opair x y) F)"]) apply (simp add: Ext Lambda_def Repl Domain) apply (simp add: Ext[symmetric]) @@ -313,7 +313,7 @@ apply (auto simp add: Fst Snd) done -lemma Elem_Lambda_Fun: "Elem (Lambda A f) (Fun U V) = (A = U & (! x. Elem x A \ Elem (f x) V))" +lemma Elem_Lambda_Fun: "Elem (Lambda A f) (Fun U V) = (A = U \ (\x. Elem x A \ Elem (f x) V))" proof - have "Elem (Lambda A f) (Fun U V) \ A = U" by (simp add: Fun_def Sep domain_Lambda) @@ -359,8 +359,8 @@ proof assume not_empty: "?Z \ Empty" note thereis_x = Regularity[where A="?Z", simplified not_empty, simplified] - then obtain x where x_def: "Elem x ?Z & (! y. Elem y x \ Not (Elem y ?Z))" .. - then have x_induct:"! y. Elem y x \ Elem y U \ P y" by (simp add: Sep) + then obtain x where x_def: "Elem x ?Z \ (\y. Elem y x \ Not (Elem y ?Z))" .. + then have x_induct:"\y. Elem y x \ Elem y U \ P y" by (simp add: Sep) have "Elem x U \ P x" by (rule impE[OF spec[OF P_induct, where x=x], OF x_induct], assumption) moreover have "Elem x U & Not(P x)" @@ -377,14 +377,14 @@ lemma cond2_wf_Elem: assumes - special_P: "? U. ! x. Not(Elem x U) \ (P x)" + special_P: "\U. \x. Not(Elem x U) \ (P x)" and P_induct: "\x. (\y. Elem y x \ P y) \ P x" shows "P a" proof - - have "? U Q. P = (\ x. (Elem x U \ Q x))" + have "\U Q. P = (\ x. (Elem x U \ Q x))" proof - - from special_P obtain U where U:"! x. Not(Elem x U) \ (P x)" .. + from special_P obtain U where U: "\x. Not(Elem x U) \ (P x)" .. show ?thesis apply (rule_tac exI[where x=U]) apply (rule exI[where x="P"]) @@ -392,7 +392,7 @@ apply (auto simp add: U) done qed - then obtain U where "? Q. P = (\ x. (Elem x U \ Q x))" .. + then obtain U where "\Q. P = (\ x. (Elem x U \ Q x))" .. then obtain Q where UQ: "P = (\ x. (Elem x U \ Q x))" .. show ?thesis apply (auto simp add: UQ) @@ -415,7 +415,7 @@ done definition Nat :: ZF - where "Nat == Sep Inf (\ N. ? n. nat2Nat n = N)" + where "Nat == Sep Inf (\N. \n. nat2Nat n = N)" lemma Elem_nat2Nat_Nat[intro]: "Elem (nat2Nat n) Nat" by (auto simp add: Nat_def Sep) @@ -427,11 +427,11 @@ by (auto simp add: Nat_def Sep Infinity) lemma no_infinite_Elem_down_chain: - "Not (? f. isFun f & Domain f = Nat & (! N. Elem N Nat \ Elem (f\(SucNat N)) (f\N)))" + "Not (\f. isFun f \ Domain f = Nat \ (\N. Elem N Nat \ Elem (f\(SucNat N)) (f\N)))" proof - { fix f - assume f:"isFun f & Domain f = Nat & (! N. Elem N Nat \ Elem (f\(SucNat N)) (f\N))" + assume f: "isFun f \ Domain f = Nat \ (\N. Elem N Nat \ Elem (f\(SucNat N)) (f\N))" let ?r = "Range f" have "?r \ Empty" apply (auto simp add: Ext Empty) @@ -439,14 +439,14 @@ apply (rule fun_value_in_range) apply (auto simp add: f Elem_Empty_Nat) done - then have "? x. Elem x ?r & (! y. Elem y x \ Not(Elem y ?r))" + then have "\x. Elem x ?r \ (\y. Elem y x \ Not(Elem y ?r))" by (simp add: Regularity) - then obtain x where x: "Elem x ?r & (! y. Elem y x \ Not(Elem y ?r))" .. - then have "? N. Elem N (Domain f) & f\N = x" + then obtain x where x: "Elem x ?r \ (\y. Elem y x \ Not(Elem y ?r))" .. + then have "\N. Elem N (Domain f) & f\N = x" apply (rule_tac fun_range_witness) apply (simp_all add: f) done - then have "? N. Elem N Nat & f\N = x" + then have "\N. Elem N Nat & f\N = x" by (simp add: f) then obtain N where N: "Elem N Nat & f\N = x" .. from N have N': "Elem N Nat" by auto @@ -475,9 +475,9 @@ assume ba: "Elem b a" let ?Z = "Upair a b" have "?Z \ Empty" by (simp add: Upair_nonEmpty) - then have "? x. Elem x ?Z & (! y. Elem y x \ Not(Elem y ?Z))" + then have "\x. Elem x ?Z \ (\y. Elem y x \ Not(Elem y ?Z))" by (simp add: Regularity) - then obtain x where x:"Elem x ?Z & (! y. Elem y x \ Not(Elem y ?Z))" .. + then obtain x where x:"Elem x ?Z \ (\y. Elem y x \ Not(Elem y ?Z))" .. then have "x = a \ x = b" by (simp add: Upair) moreover have "x = a \ Not (Elem b ?Z)" by (auto simp add: x ba) @@ -501,7 +501,7 @@ "NatInterval n 0 = Singleton (nat2Nat n)" | "NatInterval n (Suc m) = union (NatInterval n m) (Singleton (nat2Nat (n+m+1)))" -lemma n_Elem_NatInterval[rule_format]: "! q. q <= m \ Elem (nat2Nat (n+q)) (NatInterval n m)" +lemma n_Elem_NatInterval[rule_format]: "\q. q \ m \ Elem (nat2Nat (n+q)) (NatInterval n m)" apply (induct m) apply (auto simp add: Singleton union) apply (case_tac "q <= m") @@ -514,13 +514,13 @@ by (auto intro: n_Elem_NatInterval[where q = 0, simplified] simp add: Empty Ext) lemma increasing_nat2Nat[rule_format]: "0 < n \ Elem (nat2Nat (n - 1)) (nat2Nat n)" - apply (case_tac "? m. n = Suc m") + apply (case_tac "\m. n = Suc m") apply (auto simp add: SucNat_def union Singleton) apply (drule spec[where x="n - 1"]) apply arith done -lemma represent_NatInterval[rule_format]: "Elem x (NatInterval n m) \ (? u. n \ u & u \ n+m & nat2Nat u = x)" +lemma represent_NatInterval[rule_format]: "Elem x (NatInterval n m) \ (\u. n \ u \ u \ n+m \ nat2Nat u = x)" apply (induct m) apply (auto simp add: Singleton union) apply (rule_tac x="Suc (n+m)" in exI) @@ -535,12 +535,12 @@ assume mg0: "0 < m" let ?Z = "NatInterval n m" have "?Z \ Empty" by (simp add: NatInterval_not_Empty) - then have "? x. (Elem x ?Z) & (! y. Elem y x \ Not (Elem y ?Z))" + then have "\x. (Elem x ?Z) \ (\y. Elem y x \ Not (Elem y ?Z))" by (auto simp add: Regularity) - then obtain x where x:"Elem x ?Z & (! y. Elem y x \ Not (Elem y ?Z))" .. - then have "? u. n \ u & u \ n+m & nat2Nat u = x" + then obtain x where x:"Elem x ?Z \ (\y. Elem y x \ Not (Elem y ?Z))" .. + then have "\u. n \ u & u \ n+m & nat2Nat u = x" by (simp add: represent_NatInterval) - then obtain u where u: "n \ u & u \ n+m & nat2Nat u = x" .. + then obtain u where u: "n \ u & u \ n+m \ nat2Nat u = x" .. have "n < u \ False" proof assume n_less_u: "n < u" @@ -590,14 +590,14 @@ apply (case_tac "x = y") apply auto apply (case_tac "x < y") - apply (case_tac "? m. y = x + m & 0 < m") + apply (case_tac "\m. y = x + m & 0 < m") apply (auto intro: lemma_nat2Nat) apply (case_tac "y < x") - apply (case_tac "? m. x = y + m & 0 < m") + apply (case_tac "\m. x = y + m & 0 < m") apply simp apply simp using th apply blast - apply (case_tac "? m. x = y + m") + apply (case_tac "\m. x = y + m") apply (auto intro: lemma_nat2Nat) apply (drule sym) using lemma_nat2Nat apply blast @@ -625,7 +625,7 @@ (*lemma Elem_induct: "(\x. \y. Elem y x \ P y \ P x) \ P a" by (erule wf_induct[OF wf_is_Elem_of, simplified is_Elem_of_def, simplified])*) -lemma Elem_Opair_exists: "? z. Elem x z & Elem y z & Elem z (Opair x y)" +lemma Elem_Opair_exists: "\z. Elem x z & Elem y z & Elem z (Opair x y)" apply (rule exI[where x="Upair x y"]) by (simp add: Upair Opair_def) @@ -682,13 +682,13 @@ by (simp add: implode_def inj_explode) definition regular :: "(ZF * ZF) set \ bool" where - "regular R == ! A. A \ Empty \ (? x. Elem x A & (! y. (y, x) \ R \ Not (Elem y A)))" + "regular R == \A. A \ Empty \ (\x. Elem x A \ (\y. (y, x) \ R \ Not (Elem y A)))" definition set_like :: "(ZF * ZF) set \ bool" where - "set_like R == ! y. Ext R y \ range explode" + "set_like R == \y. Ext R y \ range explode" definition wfzf :: "(ZF * ZF) set \ bool" where - "wfzf R == regular R & set_like R" + "wfzf R == regular R \ set_like R" lemma regular_Elem: "regular is_Elem_of" by (simp add: regular_def is_Elem_of_def Regularity) @@ -702,7 +702,7 @@ definition SeqSum :: "(nat \ ZF) \ ZF" where "SeqSum f == Sum (Repl Nat (f o Nat2nat))" -lemma SeqSum: "Elem x (SeqSum f) = (? n. Elem x (f n))" +lemma SeqSum: "Elem x (SeqSum f) = (\n. Elem x (f n))" apply (auto simp add: SeqSum_def Sum Repl) apply (rule_tac x = "f n" in exI) apply auto @@ -732,7 +732,7 @@ lemma Elem_Ext_ZF_hull: assumes set_like_R: "set_like R" - shows "Elem x (Ext_ZF_hull R S) = (? n. Elem x (Ext_ZF_n R S n))" + shows "Elem x (Ext_ZF_hull R S) = (\n. Elem x (Ext_ZF_n R S n))" by (simp add: Ext_ZF_hull_def SeqSum) lemma Elem_Elem_Ext_ZF_hull: @@ -742,7 +742,7 @@ shows "Elem y (Ext_ZF_hull R S)" proof - from Elem_Ext_ZF_hull[OF set_like_R] x_hull - have "? n. Elem x (Ext_ZF_n R S n)" by auto + have "\n. Elem x (Ext_ZF_n R S n)" by auto then obtain n where n:"Elem x (Ext_ZF_n R S n)" .. with y_R_x have "Elem y (Ext_ZF_n R S (Suc n))" apply (auto simp add: Repl Sum) @@ -805,16 +805,16 @@ done with x show "False" by auto qed - then have "! x. P x" by auto + then have "\x. P x" by auto } - then show "(\x. (\y. (y, x) \ R \ P y) \ P x) \ (! x. P x)" by blast + then show "(\x. (\y. (y, x) \ R \ P y) \ P x) \ (\x. P x)" by blast qed lemma wf_is_Elem_of: "wf is_Elem_of" by (auto simp add: wfzf_is_Elem_of wfzf_implies_wf) lemma in_Ext_RTrans_implies_Elem_Ext_ZF_hull: - "set_like R \ x \ (Ext (R^+) s) \ Elem x (Ext_ZF_hull R s)" + "set_like R \ x \ (Ext (R\<^sup>+) s) \ Elem x (Ext_ZF_hull R s)" apply (simp add: Ext_def Elem_Ext_ZF_hull) apply (erule converse_trancl_induct[where r="R"]) apply (rule exI[where x=0]) @@ -826,21 +826,21 @@ apply (auto simp add: Elem_Ext_ZF) done -lemma implodeable_Ext_trancl: "set_like R \ set_like (R^+)" +lemma implodeable_Ext_trancl: "set_like R \ set_like (R\<^sup>+)" apply (subst set_like_def) apply (auto simp add: image_def) - apply (rule_tac x="Sep (Ext_ZF_hull R y) (\ z. z \ (Ext (R^+) y))" in exI) + apply (rule_tac x="Sep (Ext_ZF_hull R y) (\ z. z \ (Ext (R\<^sup>+) y))" in exI) apply (auto simp add: explode_def Sep set_eqI in_Ext_RTrans_implies_Elem_Ext_ZF_hull) done lemma Elem_Ext_ZF_hull_implies_in_Ext_RTrans[rule_format]: - "set_like R \ ! x. Elem x (Ext_ZF_n R s n) \ x \ (Ext (R^+) s)" + "set_like R \ \x. Elem x (Ext_ZF_n R s n) \ x \ (Ext (R\<^sup>+) s)" apply (induct_tac n) apply (auto simp add: Elem_Ext_ZF Ext_def Sum Repl) done -lemma "set_like R \ Ext_ZF (R^+) s = Ext_ZF_hull R s" +lemma "set_like R \ Ext_ZF (R\<^sup>+) s = Ext_ZF_hull R s" apply (frule implodeable_Ext_trancl) apply (auto simp add: Ext) apply (erule in_Ext_RTrans_implies_Elem_Ext_ZF_hull) @@ -856,7 +856,7 @@ show "A \ Empty \ (\x. Elem x A \ (\y. (y, x) \ R \ \ Elem y A))" proof assume A: "A \ Empty" - then have "? x. x \ explode A" + then have "\x. x \ explode A" by (auto simp add: explode_def Ext Empty) then obtain x where x:"x \ explode A" .. from iffD1[OF wf_eq_minimal wf, rule_format, where Q="explode A", OF x] @@ -873,7 +873,7 @@ apply (auto simp add: wfzf_def wf_implies_regular) done -lemma wfzf_trancl: "wfzf R \ wfzf (R^+)" +lemma wfzf_trancl: "wfzf R \ wfzf (R\<^sup>+)" by (auto simp add: wf_eq_wfzf[symmetric] implodeable_Ext_trancl wf_trancl) lemma Ext_subset_mono: "R \ S \ Ext R y \ Ext S y" diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ZF/LProd.thy --- a/src/HOL/ZF/LProd.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ZF/LProd.thy Thu Feb 15 13:04:36 2018 +0100 @@ -86,9 +86,9 @@ assumes wf_R: "wf R" shows "wf (lprod R)" proof - - have subset: "lprod (R^+) \ inv_image (mult (R^+)) mset" + have subset: "lprod (R\<^sup>+) \ inv_image (mult (R\<^sup>+)) mset" by (auto simp add: lprod_implies_mult trans_trancl) - note lprodtrancl = wf_subset[OF wf_inv_image[where r="mult (R^+)" and f="mset", + note lprodtrancl = wf_subset[OF wf_inv_image[where r="mult (R\<^sup>+)" and f="mset", OF wf_mult[OF wf_trancl[OF wf_R]]], OF subset] note lprod = wf_subset[OF lprodtrancl, where p="lprod R", OF lprod_subset, simplified] show ?thesis by (auto intro: lprod) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ZF/Zet.thy --- a/src/HOL/ZF/Zet.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ZF/Zet.thy Thu Feb 15 13:04:36 2018 +0100 @@ -17,7 +17,7 @@ definition zin :: "'a \ 'a zet \ bool" where "zin x A == x \ (Rep_zet A)" -lemma zet_ext_eq: "(A = B) = (! x. zin x A = zin x B)" +lemma zet_ext_eq: "(A = B) = (\x. zin x A = zin x B)" by (auto simp add: Rep_zet_inject[symmetric] zin_def) definition zimage :: "('a \ 'b) \ 'a zet \ 'b zet" where @@ -32,7 +32,7 @@ apply (auto simp add: explode_def Sep) done -lemma image_zet_rep: "A \ zet \ ? z . g ` A = explode z" +lemma image_zet_rep: "A \ zet \ \z . g ` A = explode z" apply (auto simp add: zet_def') apply (rule_tac x="Repl z (g o (inv_into A f))" in exI) apply (simp add: explode_Repl_eq) @@ -44,7 +44,7 @@ assumes Azet: "A \ zet" shows "g ` A \ zet" proof - - from Azet have "? (f :: _ \ ZF). inj_on f A" + from Azet have "\(f :: _ \ ZF). inj_on f A" by (auto simp add: zet_def') then obtain f where injf: "inj_on (f :: _ \ ZF) A" by auto @@ -70,7 +70,7 @@ apply (simp_all add: Rep_zet zet_image_mem) done -lemma zimage_iff: "zin y (zimage f A) = (? x. zin x A & y = f x)" +lemma zimage_iff: "zin y (zimage f A) = (\x. zin x A \ y = f x)" by (auto simp add: zin_def Rep_zimage_eq) definition zimplode :: "ZF zet \ ZF" where @@ -79,7 +79,7 @@ definition zexplode :: "ZF \ ZF zet" where "zexplode z == Abs_zet (explode z)" -lemma Rep_zet_eq_explode: "? z. Rep_zet A = explode z" +lemma Rep_zet_eq_explode: "\z. Rep_zet A = explode z" by (rule image_zet_rep[where g="\ x. x",OF Rep_zet, simplified]) lemma zexplode_zimplode: "zexplode (zimplode A) = A" @@ -117,7 +117,7 @@ "zunion a b \ Abs_zet ((Rep_zet a) \ (Rep_zet b))" definition zsubset :: "'a zet \ 'a zet \ bool" where - "zsubset a b \ ! x. zin x a \ zin x b" + "zsubset a b \ \x. zin x a \ zin x b" lemma explode_union: "explode (union a b) = (explode a) \ (explode b)" apply (rule set_eqI) @@ -126,13 +126,13 @@ lemma Rep_zet_zunion: "Rep_zet (zunion a b) = (Rep_zet a) \ (Rep_zet b)" proof - - from Rep_zet[of a] have "? f z. inj_on f (Rep_zet a) \ f ` (Rep_zet a) = explode z" + from Rep_zet[of a] have "\f z. inj_on f (Rep_zet a) \ f ` (Rep_zet a) = explode z" by (auto simp add: zet_def') then obtain fa za where a:"inj_on fa (Rep_zet a) \ fa ` (Rep_zet a) = explode za" by blast from a have fa: "inj_on fa (Rep_zet a)" by blast from a have za: "fa ` (Rep_zet a) = explode za" by blast - from Rep_zet[of b] have "? f z. inj_on f (Rep_zet b) \ f ` (Rep_zet b) = explode z" + from Rep_zet[of b] have "\f z. inj_on f (Rep_zet b) \ f ` (Rep_zet b) = explode z" by (auto simp add: zet_def') then obtain fb zb where b:"inj_on fb (Rep_zet b) \ fb ` (Rep_zet b) = explode zb" by blast diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/Zorn.thy --- a/src/HOL/Zorn.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/Zorn.thy Thu Feb 15 13:04:36 2018 +0100 @@ -669,7 +669,7 @@ from that have Ris: "R \ Chains init_seg_of" using mono_Chains [OF I_init] by blast have subch: "chain\<^sub>\ R" - using \R : Chains I\ I_init by (auto simp: init_seg_of_def chain_subset_def Chains_def) + using \R \ Chains I\ I_init by (auto simp: init_seg_of_def chain_subset_def Chains_def) have "\r\R. Refl r" and "\r\R. trans r" and "\r\R. antisym r" and "\r\R. Total r" and "\r\R. wf (r - Id)" using Chains_wo [OF \R \ Chains I\] by (simp_all add: order_on_defs) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Birthday_Paradox.thy --- a/src/HOL/ex/Birthday_Paradox.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Birthday_Paradox.thy Thu Feb 15 13:04:36 2018 +0100 @@ -36,10 +36,10 @@ } note finite_delete = this from insert have hyps: "\y \ T. card ({g. g \ extensional_funcset S (T - {y}) \ inj_on g S}) = fact (card T - 1) div fact ((card T - 1) - card S)"(is "\ _ \ T. _ = ?k") by auto from extensional_funcset_extend_domain_inj_on_eq[OF \x \ S\] - have "card {f. f : extensional_funcset (insert x S) T & inj_on f (insert x S)} = - card ((%(y, g). g(x := y)) ` {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S})" + have "card {f. f \ extensional_funcset (insert x S) T \ inj_on f (insert x S)} = + card ((\(y, g). g(x := y)) ` {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S})" by metis - also from extensional_funcset_extend_domain_inj_onI[OF \x \ S\, of T] have "... = card {(y, g). y : T & g : extensional_funcset S (T - {y}) & inj_on g S}" + also from extensional_funcset_extend_domain_inj_onI[OF \x \ S\, of T] have "\ = card {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S}" by (simp add: card_image) also have "card {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S} = card {(y, g). y \ T \ g \ {f \ extensional_funcset S (T - {y}). inj_on f S}}" by auto diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Dedekind_Real.thy --- a/src/HOL/ex/Dedekind_Real.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Dedekind_Real.thy Thu Feb 15 13:04:36 2018 +0100 @@ -1273,7 +1273,7 @@ declare equiv_realrel_iff [simp] -lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real" +lemma realrel_in_real [simp]: "realrel``{(x,y)} \ Real" by (simp add: Real_def realrel_def quotient_def, blast) declare Abs_Real_inject [simp] @@ -1283,7 +1283,7 @@ text\Case analysis on the representation of a real number as an equivalence class of pairs of positive reals.\ lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: - "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P" + "(\x y. z = Abs_Real(realrel``{(x,y)}) \ P) \ P" apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE]) apply (drule arg_cong [where f=Abs_Real]) apply (auto simp add: Rep_Real_inverse) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Executable_Relation.thy --- a/src/HOL/ex/Executable_Relation.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Executable_Relation.thy Thu Feb 15 13:04:36 2018 +0100 @@ -34,11 +34,11 @@ code_datatype Rel lemma [code]: - "member (x, y) (Rel X R) = ((x = y \ x : X) \ (x, y) : R)" + "member (x, y) (Rel X R) = ((x = y \ x \ X) \ (x, y) \ R)" by transfer auto lemma [code]: - "converse (Rel X R) = Rel X (R^-1)" + "converse (Rel X R) = Rel X (R\)" by transfer auto lemma [code]: @@ -46,11 +46,11 @@ by transfer auto lemma [code]: - "relcomp (Rel X R) (Rel Y S) = Rel (X Int Y) (Set.filter (%(x, y). y : Y) R Un (Set.filter (%(x, y). x : X) S Un R O S))" + "relcomp (Rel X R) (Rel Y S) = Rel (X \ Y) (Set.filter (\(x, y). y \ Y) R \ (Set.filter (\(x, y). x \ X) S \ R O S))" by transfer (auto simp add: Id_on_eqI relcomp.simps) lemma [code]: - "rtrancl (Rel X R) = Rel UNIV (R^+)" + "rtrancl (Rel X R) = Rel UNIV (R\<^sup>+)" apply transfer apply auto apply (metis Id_on_iff Un_commute UNIV_I rtrancl_Un_separatorE rtrancl_eq_or_trancl) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Groebner_Examples.thy --- a/src/HOL/ex/Groebner_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Groebner_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -110,7 +110,7 @@ using assms by (algebra add: collinear_def split_def fst_conv snd_conv) -lemma "EX (d::int). a*y - a*x = n*d \ EX u v. a*u + n*v = 1 \ EX e. y - x = n*e" +lemma "\(d::int). a*y - a*x = n*d \ \u v. a*u + n*v = 1 \ \e. y - x = n*e" by algebra end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Intuitionistic.thy --- a/src/HOL/ex/Intuitionistic.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Intuitionistic.thy Thu Feb 15 13:04:36 2018 +0100 @@ -86,12 +86,12 @@ ***) (*Problem 1.1*) -lemma "(ALL x. EX y. ALL z. p(x) & q(y) & r(z)) = - (ALL z. EX y. ALL x. p(x) & q(y) & r(z))" +lemma "(\x. \y. \z. p(x) \ q(y) \ r(z)) = + (\z. \y. \x. p(x) \ q(y) \ r(z))" by (iprover del: allE elim 2: allE') (*Problem 3.1*) -lemma "~ (EX x. ALL y. p y x = (~ p x x))" +lemma "\ (\x. \y. p y x = (\ p x x))" by iprover @@ -177,19 +177,19 @@ (* The converse is classical in the following implications... *) -lemma "(EX x. P(x)-->Q) --> (ALL x. P(x)) --> Q" +lemma "(\x. P(x)\Q) \ (\x. P(x)) \ Q" by iprover -lemma "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)" +lemma "((\x. P(x))\Q) \ \ (\x. P(x) \ \Q)" by iprover -lemma "((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))" +lemma "((\x. \P(x))\Q) \ \ (\x. \ (P(x)\Q))" by iprover -lemma "(ALL x. P(x)) | Q --> (ALL x. P(x) | Q)" +lemma "(\x. P(x)) \ Q \ (\x. P(x) \ Q)" by iprover -lemma "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))" +lemma "(\x. P \ Q(x)) \ (P \ (\x. Q(x)))" by iprover @@ -199,130 +199,130 @@ Some will require quantifier duplication -- not currently available*) (* Problem ~~19 *) -lemma "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))" +lemma "\\(\x. \y z. (P(y)\Q(z)) \ (P(x)\Q(x)))" by iprover (* Problem 20 *) -lemma "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w))) - --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))" +lemma "(\x y. \z. \w. (P(x)\Q(y)\R(z)\S(w))) + \ (\x y. P(x) \ Q(y)) \ (\z. R(z))" by iprover (* Problem 21 *) -lemma "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P=Q(x))" +lemma "(\x. P\Q(x)) \ (\x. Q(x)\P) \ \\(\x. P=Q(x))" by iprover (* Problem 22 *) -lemma "(ALL x. P = Q(x)) --> (P = (ALL x. Q(x)))" +lemma "(\x. P = Q(x)) \ (P = (\x. Q(x)))" by iprover (* Problem ~~23 *) -lemma "~~ ((ALL x. P | Q(x)) = (P | (ALL x. Q(x))))" +lemma "\\ ((\x. P \ Q(x)) = (P \ (\x. Q(x))))" by iprover (* Problem 25 *) -lemma "(EX x. P(x)) & - (ALL x. L(x) --> ~ (M(x) & R(x))) & - (ALL x. P(x) --> (M(x) & L(x))) & - ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) - --> (EX x. Q(x)&P(x))" +lemma "(\x. P(x)) \ + (\x. L(x) \ \ (M(x) \ R(x))) \ + (\x. P(x) \ (M(x) \ L(x))) \ + ((\x. P(x)\Q(x)) \ (\x. P(x)\R(x))) + \ (\x. Q(x)\P(x))" by iprover (* Problem 27 *) -lemma "(EX x. P(x) & ~Q(x)) & - (ALL x. P(x) --> R(x)) & - (ALL x. M(x) & L(x) --> P(x)) & - ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) - --> (ALL x. M(x) --> ~L(x))" +lemma "(\x. P(x) \ \Q(x)) \ + (\x. P(x) \ R(x)) \ + (\x. M(x) \ L(x) \ P(x)) \ + ((\x. R(x) \ \ Q(x)) \ (\x. L(x) \ \ R(x))) + \ (\x. M(x) \ \L(x))" by iprover (* Problem ~~28. AMENDED *) -lemma "(ALL x. P(x) --> (ALL x. Q(x))) & - (~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & - (~~(EX x. S(x)) --> (ALL x. L(x) --> M(x))) - --> (ALL x. P(x) & L(x) --> M(x))" +lemma "(\x. P(x) \ (\x. Q(x))) \ + (\\(\x. Q(x)\R(x)) \ (\x. Q(x)&S(x))) \ + (\\(\x. S(x)) \ (\x. L(x) \ M(x))) + \ (\x. P(x) \ L(x) \ M(x))" by iprover (* Problem 29. Essentially the same as Principia Mathematica *11.71 *) -lemma "(((EX x. P(x)) & (EX y. Q(y))) --> - (((ALL x. (P(x) --> R(x))) & (ALL y. (Q(y) --> S(y)))) = - (ALL x y. ((P(x) & Q(y)) --> (R(x) & S(y))))))" +lemma "(((\x. P(x)) \ (\y. Q(y))) \ + (((\x. (P(x) \ R(x))) \ (\y. (Q(y) \ S(y)))) = + (\x y. ((P(x) \ Q(y)) \ (R(x) \ S(y))))))" by iprover (* Problem ~~30 *) -lemma "(ALL x. (P(x) | Q(x)) --> ~ R(x)) & - (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) - --> (ALL x. ~~S(x))" +lemma "(\x. (P(x) \ Q(x)) \ \ R(x)) \ + (\x. (Q(x) \ \ S(x)) \ P(x) \ R(x)) + \ (\x. \\S(x))" by iprover (* Problem 31 *) -lemma "~(EX x. P(x) & (Q(x) | R(x))) & - (EX x. L(x) & P(x)) & - (ALL x. ~ R(x) --> M(x)) - --> (EX x. L(x) & M(x))" +lemma "\(\x. P(x) \ (Q(x) \ R(x))) \ + (\x. L(x) \ P(x)) \ + (\x. \ R(x) \ M(x)) + \ (\x. L(x) \ M(x))" by iprover (* Problem 32 *) -lemma "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) & - (ALL x. S(x) & R(x) --> L(x)) & - (ALL x. M(x) --> R(x)) - --> (ALL x. P(x) & M(x) --> L(x))" +lemma "(\x. P(x) \ (Q(x)|R(x))\S(x)) \ + (\x. S(x) \ R(x) \ L(x)) \ + (\x. M(x) \ R(x)) + \ (\x. P(x) \ M(x) \ L(x))" by iprover (* Problem ~~33 *) -lemma "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c))) = - (ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))" +lemma "(\x. \\(P(a) \ (P(x)\P(b))\P(c))) = + (\x. \\((\P(a) \ P(x) \ P(c)) \ (\P(a) \ \P(b) \ P(c))))" oops (* Problem 36 *) lemma - "(ALL x. EX y. J x y) & - (ALL x. EX y. G x y) & - (ALL x y. J x y | G x y --> (ALL z. J y z | G y z --> H x z)) - --> (ALL x. EX y. H x y)" + "(\x. \y. J x y) \ + (\x. \y. G x y) \ + (\x y. J x y \ G x y \ (\z. J y z \ G y z \ H x z)) + \ (\x. \y. H x y)" by iprover (* Problem 39 *) -lemma "~ (EX x. ALL y. F y x = (~F y y))" +lemma "\ (\x. \y. F y x = (\F y y))" by iprover (* Problem 40. AMENDED *) -lemma "(EX y. ALL x. F x y = F x x) --> - ~(ALL x. EX y. ALL z. F z y = (~ F z x))" +lemma "(\y. \x. F x y = F x x) \ + \(\x. \y. \z. F z y = (\ F z x))" by iprover (* Problem 44 *) -lemma "(ALL x. f(x) --> - (EX y. g(y) & h x y & (EX y. g(y) & ~ h x y))) & - (EX x. j(x) & (ALL y. g(y) --> h x y)) - --> (EX x. j(x) & ~f(x))" +lemma "(\x. f(x) \ + (\y. g(y) \ h x y \ (\y. g(y) \ ~ h x y))) \ + (\x. j(x) \ (\y. g(y) \ h x y)) + \ (\x. j(x) \ \f(x))" by iprover (* Problem 48 *) -lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c" +lemma "(a=b \ c=d) \ (a=c \ b=d) \ a=d \ b=c" by iprover (* Problem 51 *) -lemma "((EX z w. (ALL x y. (P x y = ((x = z) & (y = w))))) --> - (EX z. (ALL x. (EX w. ((ALL y. (P x y = (y = w))) = (x = z))))))" +lemma "((\z w. (\x y. (P x y = ((x = z) \ (y = w))))) \ + (\z. (\x. (\w. ((\y. (P x y = (y = w))) = (x = z))))))" by iprover (* Problem 52 *) (*Almost the same as 51. *) -lemma "((EX z w. (ALL x y. (P x y = ((x = z) & (y = w))))) --> - (EX w. (ALL y. (EX z. ((ALL x. (P x y = (x = z))) = (y = w))))))" +lemma "((\z w. (\x y. (P x y = ((x = z) \ (y = w))))) \ + (\w. (\y. (\z. ((\x. (P x y = (x = z))) = (y = w))))))" by iprover (* Problem 56 *) -lemma "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) = (ALL x. P(x) --> P(f(x)))" +lemma "(\x. (\y. P(y) \ x=f(y)) \ P(x)) = (\x. P(x) \ P(f(x)))" by iprover (* Problem 57 *) -lemma "P (f a b) (f b c) & P (f b c) (f a c) & - (ALL x y z. P x y & P y z --> P x z) --> P (f a b) (f a c)" +lemma "P (f a b) (f b c) & P (f b c) (f a c) \ + (\x y z. P x y \ P y z \ P x z) \ P (f a b) (f a c)" by iprover (* Problem 60 *) -lemma "ALL x. P x (f x) = (EX y. (ALL z. P z y --> P z (f x)) & P x y)" +lemma "\x. P x (f x) = (\y. (\z. P z y \ P z (f x)) \ P x y)" by iprover end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/LocaleTest2.thy --- a/src/HOL/ex/LocaleTest2.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/LocaleTest2.thy Thu Feb 15 13:04:36 2018 +0100 @@ -43,7 +43,7 @@ where "(x \ y) = (x \ y & x ~= y)" theorem abs_test: - "(\) = (%x y. x \ y)" + "(\) = (\x y. x \ y)" by simp definition @@ -57,8 +57,8 @@ end locale dlat = dpo + - assumes ex_inf: "EX inf. dpo.is_inf le x y inf" - and ex_sup: "EX sup. dpo.is_sup le x y sup" + assumes ex_inf: "\inf. dpo.is_inf le x y inf" + and ex_sup: "\sup. dpo.is_sup le x y sup" begin @@ -436,11 +436,11 @@ proof fix x y from total have "is_inf x y (if x \ y then x else y)" by (auto simp: is_inf_def) - then show "EX inf. is_inf x y inf" by blast + then show "\inf. is_inf x y inf" by blast next fix x y from total have "is_sup x y (if x \ y then y else x)" by (auto simp: is_sup_def) - then show "EX sup. is_sup x y sup" by blast + then show "\sup. is_sup x y sup" by blast qed sublocale dlo < ddlat @@ -644,7 +644,7 @@ definition inv where "inv x = (THE y. x ** y = one & y ** x = one)" definition - unit where "unit x = (EX y. x ** y = one & y ** x = one)" + unit where "unit x = (\y. x ** y = one & y ** x = one)" lemma inv_unique: assumes eq: "y ** x = one" "x ** y' = one" @@ -872,7 +872,7 @@ then have inv: "f o Hilbert_Choice.inv f = id & Hilbert_Choice.inv f o f = id" by (simp add: bij_def surj_iff inj_iff) - show "EX g. f o g = id & g o f = id" by rule (rule inv) + show "\g. f \ g = id \ g \ f = id" by rule (rule inv) qed qed @@ -896,7 +896,7 @@ apply (insert unit_id) apply simp done - show "Dmonoid.inv (o) id f = inv (f :: unit => unit)" + show "Dmonoid.inv (o) id f = inv (f :: unit \ unit)" apply (unfold Dmonoid.inv_def [OF Dmonoid]) apply (insert unit_id) apply simp diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Refute_Examples.thy --- a/src/HOL/ex/Refute_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Refute_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -247,13 +247,13 @@ text \"The transitive closure 'T' of an arbitrary relation 'P' is non-empty."\ definition "trans" :: "('a \ 'a \ bool) \ bool" where - "trans P == (ALL x y z. P x y \ P y z \ P x z)" + "trans P \ (\x y z. P x y \ P y z \ P x z)" definition "subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where - "subset P Q == (ALL x y. P x y \ Q x y)" + "subset P Q \ (\x y. P x y \ Q x y)" definition "trans_closure" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where - "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R \ trans R \ subset P R)" + "trans_closure P Q \ (subset Q P) \ (trans P) \ (\R. subset Q R \ trans R \ subset P R)" lemma "trans_closure T P \ (\x y. T x y)" refute [expect = genuine] @@ -372,15 +372,15 @@ refute by simp -lemma "x : {x. P x}" +lemma "x \ {x. P x}" refute oops -lemma "P (:)" +lemma "P (\)" refute oops -lemma "P ((:) x)" +lemma "P ((\) x)" refute oops @@ -388,11 +388,11 @@ refute oops -lemma "A Un B = A Int B" +lemma "A \ B = A \ B" refute oops -lemma "(A Int B) Un C = (A Un C) Int B" +lemma "(A \ B) \ C = (A \ C) \ B" refute oops diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/SAT_Examples.thy --- a/src/HOL/ex/SAT_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/SAT_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -76,7 +76,7 @@ text \eta-Equivalence\ -lemma "(ALL x. P x) | ~ All P" +lemma "(\x. P x) \ \ All P" by sat declare [[sat_trace = false]] diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy --- a/src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Set_Comprehension_Pointfree_Examples.thy Thu Feb 15 13:04:36 2018 +0100 @@ -12,102 +12,102 @@ declare [[simproc add: finite_Collect]] lemma - "finite (UNIV::'a set) ==> finite {p. EX x::'a. p = (x, x)}" + "finite (UNIV::'a set) \ finite {p. \x::'a. p = (x, x)}" by simp lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B}" + "finite A \ finite B \ finite {f a b| a b. a \ A \ b \ B}" by simp lemma - "finite B ==> finite A' ==> finite {f a b| a b. a : A \ a : A' \ b : B}" + "finite B \ finite A' \ finite {f a b| a b. a \ A \ a \ A' \ b \ B}" by simp lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ b : B'}" + "finite A \ finite B \ finite {f a b| a b. a \ A \ b \ B \ b \ B'}" by simp lemma - "finite A ==> finite B ==> finite C ==> finite {f a b c| a b c. a : A \ b : B \ c : C}" + "finite A \ finite B \ finite C \ finite {f a b c| a b c. a \ A \ b \ B \ c \ C}" by simp lemma - "finite A ==> finite B ==> finite C ==> finite D ==> - finite {f a b c d| a b c d. a : A \ b : B \ c : C \ d : D}" + "finite A \ finite B \ finite C \ finite D \ + finite {f a b c d| a b c d. a \ A \ b \ B \ c \ C \ d \ D}" by simp lemma - "finite A ==> finite B ==> finite C ==> finite D ==> finite E ==> - finite {f a b c d e | a b c d e. a : A \ b : B \ c : C \ d : D \ e : E}" + "finite A \ finite B \ finite C \ finite D \ finite E \ + finite {f a b c d e | a b c d e. a \ A \ b \ B \ c \ C \ d \ D \ e \ E}" by simp lemma - "finite A ==> finite B ==> finite C ==> finite D ==> finite E \ - finite {f a d c b e | e b c d a. b : B \ a : A \ e : E' \ c : C \ d : D \ e : E \ b : B'}" + "finite A \ finite B \ finite C \ finite D \ finite E \ + finite {f a d c b e | e b c d a. b \ B \ a \ A \ e \ E' \ c \ C \ d \ D \ e \ E \ b \ B'}" by simp lemma "\ finite A ; finite B ; finite C ; finite D \ - \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" + \ finite ({f a b c d| a b c d. a \ A \ b \ B \ c \ C \ d \ D})" by simp lemma "finite ((\(a,b,c,d). f a b c d) ` (A \ B \ C \ D)) - \ finite ({f a b c d| a b c d. a : A \ b : B \ c : C \ d : D})" + \ finite ({f a b c d| a b c d. a \ A \ b \ B \ c \ C \ d \ D})" by simp lemma - "finite S ==> finite {s'. EX s:S. s' = f a e s}" + "finite S \ finite {s'. \s\S. s' = f a e s}" by simp lemma - "finite A ==> finite B ==> finite {f a b| a b. a : A \ b : B \ a \ Z}" + "finite A \ finite B \ finite {f a b| a b. a \ A \ b \ B \ a \ Z}" by simp lemma - "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ b : B \ (x,y) \ R}" + "finite A \ finite B \ finite R \ finite {f a b x y| a b x y. a \ A \ b \ B \ (x,y) \ R}" by simp lemma - "finite A ==> finite B ==> finite R ==> finite {f a b x y| a b x y. a : A \ (x,y) \ R \ b : B}" + "finite A \ finite B \ finite R \ finite {f a b x y| a b x y. a \ A \ (x,y) \ R \ b \ B}" by simp lemma - "finite A ==> finite B ==> finite R ==> finite {f a (x, b) y| y b x a. a : A \ (x,y) \ R \ b : B}" + "finite A \ finite B \ finite R \ finite {f a (x, b) y| y b x a. a \ A \ (x,y) \ R \ b \ B}" by simp lemma - "finite A ==> finite AA ==> finite B ==> finite {f a b| a b. (a : A \ a : AA) \ b : B \ a \ Z}" + "finite A \ finite AA \ finite B \ finite {f a b| a b. (a \ A \ a \ AA) \ b \ B \ a \ Z}" by simp lemma - "finite A1 ==> finite A2 ==> finite A3 ==> finite A4 ==> finite A5 ==> finite B ==> - finite {f a b c | a b c. ((a : A1 \ a : A2) \ (a : A3 \ (a : A4 \ a : A5))) \ b : B \ a \ Z}" + "finite A1 \ finite A2 \ finite A3 \ finite A4 \ finite A5 \ finite B \ + finite {f a b c | a b c. ((a \ A1 \ a \ A2) \ (a \ A3 \ (a \ A4 \ a \ A5))) \ b \ B \ a \ Z}" apply simp oops -lemma "finite B ==> finite {c. EX x. x : B & c = a * x}" +lemma "finite B \ finite {c. \x. x \ B \ c = a * x}" by simp lemma - "finite A ==> finite B ==> finite {f a * g b |a b. a : A & b : B}" + "finite A \ finite B \ finite {f a * g b |a b. a \ A \ b \ B}" by simp lemma - "finite S ==> inj (%(x, y). g x y) ==> finite {f x y| x y. g x y : S}" + "finite S \ inj (\(x, y). g x y) \ finite {f x y| x y. g x y \ S}" by (auto intro: finite_vimageI) lemma - "finite A ==> finite S ==> inj (%(x, y). g x y) ==> finite {f x y z | x y z. g x y : S & z : A}" + "finite A \ finite S \ inj (\(x, y). g x y) \ finite {f x y z | x y z. g x y \ S & z \ A}" by (auto intro: finite_vimageI) lemma - "finite S ==> finite A ==> inj (%(x, y). g x y) ==> inj (%(x, y). h x y) - ==> finite {f a b c d | a b c d. g a c : S & h b d : A}" + "finite S \ finite A \ inj (\(x, y). g x y) \ inj (\(x, y). h x y) + \ finite {f a b c d | a b c d. g a c \ S \ h b d \ A}" by (auto intro: finite_vimageI) lemma - assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) : S}" + assumes "finite S" shows "finite {(a,b,c,d). ([a, b], [c, d]) \ S}" using assms by (auto intro!: finite_vimageI simp add: inj_on_def) (* injectivity to be automated with further rules and automation *) @@ -123,15 +123,15 @@ definition union :: "nat set => nat set => nat set" where - "union A B = {x. x : A \ x : B}" + "union A B = {x. x \ A \ x \ B}" -definition common_subsets :: "nat set => nat set => nat set set" +definition common_subsets :: "nat set \ nat set \ nat set set" where "common_subsets S1 S2 = {S. S \ S1 \ S \ S2}" definition products :: "nat set => nat set => nat set" where - "products A B = {c. EX a b. a : A & b : B & c = a * b}" + "products A B = {c. \a b. a \ A \ b \ B \ c = a * b}" export_code products in Haskell diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Set_Theory.thy --- a/src/HOL/ex/Set_Theory.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Set_Theory.thy Thu Feb 15 13:04:36 2018 +0100 @@ -129,12 +129,12 @@ have "finite A" using \card A \ 2\ by(auto intro:ccontr) have 0: "R `` A <= A" using \sym R\ \Domain R <= A\ unfolding Domain_unfold sym_def by blast - have h: "ALL a:A. R `` {a} <= A" using 0 by blast - hence 1: "ALL a:A. finite(R `` {a})" using \finite A\ + have h: "\a\A. R `` {a} <= A" using 0 by blast + hence 1: "\a\A. finite(R `` {a})" using \finite A\ by(blast intro: finite_subset) have sub: "?N ` A <= {0..a\A. R `` {a} - {a} < A" using h by blast thus ?thesis using psubset_card_mono[OF \finite A\] by auto qed show "~ inj_on ?N A" (is "~ ?I") @@ -143,8 +143,8 @@ hence "?n = card(?N ` A)" by(rule card_image[symmetric]) with sub \finite A\ have 2[simp]: "?N ` A = {0..card A \ 2\ by simp+ - then obtain a b where ab: "a:A" "b:A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1" + have "0 \ ?N ` A" and "?n - 1 \ ?N ` A" using \card A \ 2\ by simp+ + then obtain a b where ab: "a\A" "b\A" and Na: "?N a = 0" and Nb: "?N b = ?n - 1" by (auto simp del: 2) have "a \ b" using Na Nb \card A \ 2\ by auto have "R `` {a} - {a} = {}" by (metis 1 Na ab card_eq_0_iff finite_Diff) diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/Tarski.thy --- a/src/HOL/ex/Tarski.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/Tarski.thy Thu Feb 15 13:04:36 2018 +0100 @@ -385,7 +385,7 @@ by (simp add: isLub_def A_def r_def) lemma isLubI: - "\L \ A; \y \ S. (y, L) \ r; (\z \ A. (\y \ S. (y, z):r) \ (L, z) \ r)\ \ isLub S cl L" + "\L \ A; \y \ S. (y, L) \ r; (\z \ A. (\y \ S. (y, z)\r) \ (L, z) \ r)\ \ isLub S cl L" by (simp add: isLub_def A_def r_def) end diff -r e4e57da0583a -r 560fbd6bc047 src/HOL/ex/While_Combinator_Example.thy --- a/src/HOL/ex/While_Combinator_Example.thy Wed Feb 14 16:32:09 2018 +0100 +++ b/src/HOL/ex/While_Combinator_Example.thy Thu Feb 15 13:04:36 2018 +0100 @@ -41,9 +41,9 @@ theorem "P (lfp (\N::int set. {0} \ {(n + 2) mod 6 | n. n \ N})) = P {0, 4, 2}" proof - - have seteq: "!!A B. (A = B) = ((!a : A. a:B) & (!b:B. b:A))" + have seteq: "\A B. (A = B) = ((\a \ A. a\B) \ (\b\B. b\A))" by blast - have aux: "!!f A B. {f n | n. A n \ B n} = {f n | n. A n} \ {f n | n. B n}" + have aux: "\f A B. {f n | n. A n \ B n} = {f n | n. A n} \ {f n | n. B n}" apply blast done show ?thesis