diff -r 1f4dfcd9351b -r b89b578d15e6 Quot/Examples/Larry.thy --- a/Quot/Examples/Larry.thy Thu Dec 10 18:28:41 2009 +0100 +++ b/Quot/Examples/Larry.thy Thu Dec 10 19:05:56 2009 +0100 @@ -22,6 +22,8 @@ | SYM: "X \ Y \ Y \ X" | TRANS: "\X \ Y; Y \ Z\ \ X \ Z" +lemmas msgrel.intros[intro] + text{*Proving that it is an equivalence relation*} lemma msgrel_refl: "X \ X" @@ -38,7 +40,7 @@ subsubsection{*The Set of Nonces*} -primrec +fun freenonces :: "freemsg \ nat set" where "freenonces (NONCE N) = {N}" @@ -47,8 +49,9 @@ | "freenonces (DECRYPT K X) = freenonces X" theorem msgrel_imp_eq_freenonces: - "U \ V \ freenonces U = freenonces V" -by (erule msgrel.induct, auto) + assumes a: "U \ V" + shows "freenonces U = freenonces V" +using a by (induct) (auto) subsubsection{*The Left Projection*} @@ -67,15 +70,13 @@ (the abstract constructor) is injective*} lemma msgrel_imp_eqv_freeleft_aux: shows "freeleft U \ freeleft U" -apply(induct rule: freeleft.induct) -apply(auto intro: msgrel.intros) -done +by (induct rule: freeleft.induct) (auto) theorem msgrel_imp_eqv_freeleft: - "U \ V \ freeleft U \ freeleft V" -apply(erule msgrel.induct) -apply(auto intro: msgrel.intros msgrel_imp_eqv_freeleft_aux) -done + assumes a: "U \ V" + shows "freeleft U \ freeleft V" +using a +by (induct)(auto intro: msgrel_imp_eqv_freeleft_aux) subsubsection{*The Right Projection*} @@ -93,13 +94,13 @@ (the abstract constructor) is injective*} lemma msgrel_imp_eqv_freeright_aux: shows "freeright U \ freeright U" -apply(induct rule: freeright.induct) -apply(auto intro: msgrel.intros) -done +by (induct rule: freeright.induct) (auto) theorem msgrel_imp_eqv_freeright: - "U \ V \ freeright U \ freeright V" -by (erule msgrel.induct, auto intro: msgrel.intros msgrel_imp_eqv_freeright_aux) + assumes a: "U \ V" + shows "freeright U \ freeright V" +using a +by (induct) (auto intro: msgrel_imp_eqv_freeright_aux) subsubsection{*The Discriminator for Constructors*} @@ -114,8 +115,9 @@ text{*This theorem helps us prove @{term "Nonce N \ MPair X Y"}*} theorem msgrel_imp_eq_freediscrim: - "U \ V \ freediscrim U = freediscrim V" -by (erule msgrel.induct, auto) + assumes a: "U \ V" + shows "freediscrim U = freediscrim V" +using a by (induct, auto) subsection{*The Initial Algebra: A Quotiented Message Type*} @@ -261,14 +263,17 @@ text{*Can also be proved using the function @{term nonces}*} lemma Nonce_Nonce_eq [iff]: shows "(Nonce m = Nonce n) = (m = n)" -apply(rule iffI) -apply(lifting NONCE_imp_eq) -apply(simp) -done +proof + assume "Nonce m = Nonce n" + then show "m = n" by (lifting NONCE_imp_eq) +next + assume "m = n" + then show "Nonce m = Nonce n" by simp +qed lemma MPAIR_imp_eqv_left: shows "MPAIR X Y \ MPAIR X' Y' \ X \ X'" -by (drule msgrel_imp_eqv_freeleft, simp) +by (drule msgrel_imp_eqv_freeleft) (simp) lemma MPair_imp_eq_left: assumes eq: "MPair X Y = MPair X' Y'" @@ -277,7 +282,7 @@ lemma MPAIR_imp_eqv_right: shows "MPAIR X Y \ MPAIR X' Y' \ Y \ Y'" -by (drule msgrel_imp_eqv_freeright, simp) +by (drule msgrel_imp_eqv_freeright) (simp) lemma MPair_imp_eq_right: shows "MPair X Y = MPair X' Y' \ Y = Y'"