Adapted to new inductive definition package.
--- a/src/HOL/Algebra/FiniteProduct.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy Wed Jul 11 11:14:51 2007 +0200
@@ -18,13 +18,12 @@
@{text "x \<in> carrier G"}. We introduce an explicit argument for the domain
@{text D}. *}
-consts
+inductive_set
foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set"
-
-inductive "foldSetD D f e"
- intros
+ for D :: "'a set" and f :: "'b => 'a => 'a" and e :: 'a
+ where
emptyI [intro]: "e \<in> D ==> ({}, e) \<in> foldSetD D f e"
- insertI [intro]: "[| x ~: A; f x y \<in> D; (A, y) \<in> foldSetD D f e |] ==>
+ | insertI [intro]: "[| x ~: A; f x y \<in> D; (A, y) \<in> foldSetD D f e |] ==>
(insert x A, f x y) \<in> foldSetD D f e"
inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> foldSetD D f e"
@@ -36,7 +35,7 @@
lemma foldSetD_closed:
"[| (A, z) \<in> foldSetD D f e ; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D
|] ==> z \<in> D";
- by (erule foldSetD.elims) auto
+ by (erule foldSetD.cases) auto
lemma Diff1_foldSetD:
"[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; f x y \<in> D |] ==>
@@ -75,7 +74,7 @@
lemma (in LCD) foldSetD_closed [dest]:
"(A, z) \<in> foldSetD D f e ==> z \<in> D";
- by (erule foldSetD.elims) auto
+ by (erule foldSetD.cases) auto
lemma (in LCD) Diff1_foldSetD:
"[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; A \<subseteq> B |] ==>
@@ -117,7 +116,7 @@
apply (erule rev_mp)
apply (simp add: less_Suc_eq_le)
apply (rule impI)
- apply (rename_tac Aa xa ya Ab xb yb, case_tac "xa = xb")
+ apply (rename_tac xa Aa ya xb Ab yb, case_tac "xa = xb")
apply (subgoal_tac "Aa = Ab")
prefer 2 apply (blast elim!: equalityE)
apply blast
--- a/src/HOL/Auth/CertifiedEmail.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/CertifiedEmail.thy Wed Jul 11 11:14:51 2007 +0200
@@ -30,24 +30,23 @@
"response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}"
-consts certified_mail :: "event list set"
-inductive "certified_mail"
- intros
+inductive_set certified_mail :: "event list set"
+ where
-Nil: --{*The empty trace*}
+ Nil: --{*The empty trace*}
"[] \<in> certified_mail"
-Fake: --{*The Spy may say anything he can say. The sender field is correct,
+| Fake: --{*The Spy may say anything he can say. The sender field is correct,
but agents don't use that information.*}
"[| evsf \<in> certified_mail; X \<in> synth(analz(spies evsf))|]
==> Says Spy B X # evsf \<in> certified_mail"
-FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
+| FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
equipped with the necessary credentials to serve as an SSL server.*}
"[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
-CM1: --{*The sender approaches the recipient. The message is a number.*}
+| CM1: --{*The sender approaches the recipient. The message is a number.*}
"[|evs1 \<in> certified_mail;
Key K \<notin> used evs1;
K \<in> symKeys;
@@ -58,7 +57,7 @@
Number cleartext, Nonce q, S2TTP|} # evs1
\<in> certified_mail"
-CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
+| CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
password to @{term TTP} over an SSL channel.*}
"[|evs2 \<in> certified_mail;
Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext,
@@ -69,7 +68,7 @@
Notes TTP {|Agent R, Agent TTP, S2TTP, Key(RPwd R), hr|} # evs2
\<in> certified_mail"
-CM3: --{*@{term TTP} simultaneously reveals the key to the recipient and gives
+| CM3: --{*@{term TTP} simultaneously reveals the key to the recipient and gives
a receipt to the sender. The SSL channel does not authenticate
the client (@{term R}), but @{term TTP} accepts the message only
if the given password is that of the claimed sender, @{term R}.
@@ -84,7 +83,7 @@
Gets S (Crypt (priSK TTP) S2TTP) #
Says TTP S (Crypt (priSK TTP) S2TTP) # evs3 \<in> certified_mail"
-Reception:
+| Reception:
"[|evsr \<in> certified_mail; Says A B X \<in> set evsr|]
==> Gets B X#evsr \<in> certified_mail"
--- a/src/HOL/Auth/Guard/Analz.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Analz.thy Wed Jul 11 11:14:51 2007 +0200
@@ -18,13 +18,13 @@
subsection{*messages that do not contribute to analz*}
-consts pparts :: "msg set => msg set"
-
-inductive "pparts H"
-intros
-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"
+inductive_set
+ 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"
subsection{*basic facts about @{term pparts}*}
@@ -120,13 +120,13 @@
subsection{*messages that contribute to analz*}
-consts kparts :: "msg set => msg set"
-
-inductive "kparts H"
-intros
-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"
+inductive_set
+ 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"
subsection{*basic facts about @{term kparts}*}
--- a/src/HOL/Auth/Guard/Guard.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Guard.thy Wed Jul 11 11:14:51 2007 +0200
@@ -18,14 +18,14 @@
in a sub-message of the form Crypt (invKey K) X with K:Ks
******************************************************************************)
-consts guard :: "nat => key set => msg set"
-
-inductive "guard n Ks"
-intros
-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"
+inductive_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"
subsection{*basic facts about @{term guard}*}
@@ -117,7 +117,7 @@
==> Guard n Ks (analz G)"
apply (auto simp: Guard_def)
apply (erule analz.induct, auto)
-by (ind_cases "Crypt K Xa:guard n Ks", 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"
by (auto simp: Guard_def)
--- a/src/HOL/Auth/Guard/GuardK.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/GuardK.thy Wed Jul 11 11:14:51 2007 +0200
@@ -23,14 +23,14 @@
in a sub-message of the form Crypt (invKey K) X with K:Ks
******************************************************************************)
-consts guardK :: "nat => key set => msg set"
-
-inductive "guardK n Ks"
-intros
-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"
+inductive_set
+ 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"
subsection{*basic facts about @{term guardK}*}
@@ -119,7 +119,7 @@
==> GuardK n Ks (analz G)"
apply (auto simp: GuardK_def)
apply (erule analz.induct, auto)
-by (ind_cases "Crypt K Xa:guardK n Ks", 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"
by (auto simp: GuardK_def)
--- a/src/HOL/Auth/Guard/Guard_NS_Public.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy Wed Jul 11 11:14:51 2007 +0200
@@ -40,22 +40,20 @@
subsection{*definition of the protocol*}
-consts nsp :: "event list set"
+inductive_set nsp :: "event list set"
+where
-inductive nsp
-intros
+ 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: "[| 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*}
@@ -72,7 +70,7 @@
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", 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)"
--- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy Wed Jul 11 11:14:51 2007 +0200
@@ -62,25 +62,23 @@
subsection{*definition of the protocol*}
-consts or :: "event list set"
+inductive_set or :: "event list set"
+where
-inductive or
-intros
+ 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*}
@@ -97,7 +95,7 @@
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", 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)"
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Wed Jul 11 11:14:51 2007 +0200
@@ -53,25 +53,23 @@
subsection{*definition of the protocol*}
-consts ya :: "event list set"
+inductive_set ya :: "event list set"
+where
-inductive ya
-intros
+ 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*}
@@ -88,7 +86,7 @@
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", 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)"
--- a/src/HOL/Auth/Guard/List_Msg.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/List_Msg.thy Wed Jul 11 11:14:51 2007 +0200
@@ -137,12 +137,10 @@
nil :: msg where
"nil == Number 0"
-consts agl :: "msg set"
-
-inductive agl
-intros
-Nil[intro]: "nil:agl"
-Cons[intro]: "[| A:agent; I:agl |] ==> cons (Agent A) I :agl"
+inductive_set agl :: "msg set"
+where
+ Nil[intro]: "nil:agl"
+| Cons[intro]: "[| A:agent; I:agl |] ==> cons (Agent A) I :agl"
subsubsection{*basic facts about agent lists*}
--- a/src/HOL/Auth/Guard/P1.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/P1.thy Wed Jul 11 11:14:51 2007 +0200
@@ -150,20 +150,18 @@
subsubsection{*protocol*}
-consts p1 :: "event list set"
+inductive_set p1 :: "event list set"
+where
-inductive p1
-intros
-
-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*}
@@ -181,13 +179,13 @@
subsubsection{*Valid Offer Lists*}
-consts valid :: "agent => nat => agent => msg set"
+inductive_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"
-inductive "valid A n B"
-intros
-Request [intro]: "cons (anchor A n B) nil:valid A n B"
-
-Propose [intro]: "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*}
@@ -284,15 +282,15 @@
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")
-apply (ind_cases "{|x,M,l'a|}:valid A n B")
+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")
+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")
+apply (ind_cases "{|x,l'|}:valid A n B" for x l')
by (drule_tac x=l' in spec, simp, blast)
subsubsection{*insertion resilience:
@@ -308,15 +306,15 @@
(* i = 0 *)
apply clarify
apply (frule len_not_empty, clarsimp)
-apply (ind_cases "{|x,l'|}:valid A n B", simp)
-apply (ind_cases "{|x,M,l'|}:valid A n B", clarsimp)
-apply (ind_cases "{|head l',l'|}:valid A n B", 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")
+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")
+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)
@@ -329,14 +327,14 @@
(* i = 0 *)
apply clarify
apply (frule len_not_empty, clarsimp)
-apply (ind_cases "{|x,l'|}:valid A n B")
+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")
+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")
+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)
@@ -375,7 +373,7 @@
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", 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)"
--- a/src/HOL/Auth/Guard/P2.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/P2.thy Wed Jul 11 11:14:51 2007 +0200
@@ -130,31 +130,29 @@
subsubsection{*protocol*}
-consts p2 :: "event list set"
+inductive_set p2 :: "event list set"
+where
-inductive p2
-intros
-
-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*}
-consts valid :: "agent => nat => agent => msg set"
+inductive_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"
-inductive "valid A n B"
-intros
-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*}
@@ -188,15 +186,15 @@
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")
-apply (ind_cases "{|x,M,l'a|}:valid A n B")
+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")
+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")
+apply (ind_cases "{|x,l'|}:valid A n B" for x l')
by (drule_tac x=l' in spec, simp, blast)
subsection{*insertion resilience:
@@ -212,15 +210,15 @@
(* i = 0 *)
apply clarify
apply (frule len_not_empty, clarsimp)
-apply (ind_cases "{|x,l'|}:valid A n B", simp)
-apply (ind_cases "{|x,M,l'|}:valid A n B", clarsimp)
-apply (ind_cases "{|head l',l'|}:valid A n B", 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")
+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")
+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)
@@ -233,14 +231,14 @@
(* i = 0 *)
apply clarify
apply (frule len_not_empty, clarsimp)
-apply (ind_cases "{|x,l'|}:valid A n B")
+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")
+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")
+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)
@@ -279,7 +277,7 @@
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", 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)"
--- a/src/HOL/Auth/Guard/Proto.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Guard/Proto.thy Wed Jul 11 11:14:51 2007 +0200
@@ -106,22 +106,23 @@
"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))"
-consts tr :: "proto => event list set"
-
-inductive "tr p" intros
+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", auto)
+by (ind_cases "ev # evs:tr p" for ev evs, auto)
constdefs has_only_Says' :: "proto => bool"
"has_only_Says' p == ALL R. R:p --> is_Says (snd R)"
@@ -379,9 +380,6 @@
Na :: nat "Na == 0"
Nb :: nat "Nb == 1"
-consts
-ns :: proto
-
abbreviation
ns1 :: rule where
"ns1 == ({}, Says a b (Crypt (pubK b) {|Nonce Na, Agent a|}))"
@@ -397,10 +395,10 @@
Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})},
Says a b (Crypt (pubK b) (Nonce Nb)))"
-inductive ns intros
-[iff]: "ns1:ns"
-[iff]: "ns2:ns"
-[iff]: "ns3:ns"
+inductive_set ns :: proto where
+ [iff]: "ns1:ns"
+| [iff]: "ns2:ns"
+| [iff]: "ns3:ns"
abbreviation (input)
ns3a :: event where
@@ -428,6 +426,8 @@
lemma inf_is_ord [iff]: "ord ns inf"
apply (unfold ord_def inf_def)
apply (rule allI)+
+apply (rule impI)
+apply (simp add: split_paired_all)
by (rule impI, erule ns.cases, simp_all)+
subsection{*general properties*}
@@ -435,6 +435,7 @@
lemma ns_has_only_Says' [iff]: "has_only_Says' ns"
apply (unfold has_only_Says'_def)
apply (rule allI, rule impI)
+apply (simp add: split_paired_all)
by (erule ns.cases, auto)
lemma newn_ns1 [iff]: "newn ns1 = {Na}"
@@ -458,6 +459,7 @@
apply (erule fresh_ruleD, simp, simp, simp, simp)
apply (rule allI)+
apply (rule impI, rule impI, rule impI)
+apply (simp add: split_paired_all)
apply (erule ns.cases)
(* fresh with NS1 *)
apply (rule impI, rule impI, rule impI, rule impI, rule impI, rule impI)
@@ -525,6 +527,7 @@
lemma "uniq' ns inf secret"
apply (unfold uniq'_def)
apply (rule allI)+
+apply (simp add: split_paired_all)
apply (rule impI, erule ns.cases)
(* R = ns1 *)
apply (rule impI, erule ns.cases)
@@ -540,7 +543,8 @@
apply (drule Crypt_insert_synth, simp, simp, simp)
apply (drule Crypt_insert_synth, simp, simp, simp, simp)
(* Proto *)
-apply (erule_tac P="ok evsa Ra sa" in rev_mp)
+apply (erule_tac P="ok evsa R sa" in rev_mp)
+apply (simp add: split_paired_all)
apply (erule ns.cases)
(* ns1 *)
apply (clarify, simp add: secret_def)
@@ -563,7 +567,8 @@
apply (drule Crypt_insert_synth, simp, simp, simp)
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp, simp)
(* Proto *)
-apply (erule_tac P="ok evsa Ra sa" in rev_mp)
+apply (erule_tac P="ok evsa R sa" in rev_mp)
+apply (simp add: split_paired_all)
apply (erule ns.cases)
(* ns1 *)
apply (clarify, simp add: secret_def)
@@ -591,7 +596,8 @@
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp)
apply (drule_tac n="nonce s' Nb" in Crypt_insert_synth, simp, simp, simp, simp)
(* Proto *)
-apply (erule_tac P="ok evsa Ra sa" in rev_mp)
+apply (erule_tac P="ok evsa R sa" in rev_mp)
+apply (simp add: split_paired_all)
apply (erule ns.cases)
(* ns1 *)
apply (simp add: secret_def)
--- a/src/HOL/Auth/KerberosIV.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/KerberosIV.thy Wed Jul 11 11:14:51 2007 +0200
@@ -110,19 +110,16 @@
Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
\<in> set evs"
-consts
-
-kerbIV :: "event list set"
-inductive "kerbIV"
- intros
+inductive_set kerbIV :: "event list set"
+ where
Nil: "[] \<in> kerbIV"
- Fake: "\<lbrakk> evsf \<in> kerbIV; X \<in> synth (analz (spies evsf)) \<rbrakk>
+ | Fake: "\<lbrakk> evsf \<in> kerbIV; X \<in> synth (analz (spies evsf)) \<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> kerbIV"
(* FROM the initiator *)
- K1: "\<lbrakk> evs1 \<in> kerbIV \<rbrakk>
+ | K1: "\<lbrakk> evs1 \<in> kerbIV \<rbrakk>
\<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
\<in> kerbIV"
@@ -133,7 +130,7 @@
(*---------------------------------------------------------------------*)
(*FROM Kas *)
- K2: "\<lbrakk> evs2 \<in> kerbIV; Key authK \<notin> used evs2; authK \<in> symKeys;
+ | K2: "\<lbrakk> evs2 \<in> kerbIV; Key authK \<notin> used evs2; authK \<in> symKeys;
Says A' Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Kas A
(Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2),
@@ -149,7 +146,7 @@
(*---------------------------------------------------------------------*)
(* FROM the initiator *)
- K3: "\<lbrakk> evs3 \<in> kerbIV;
+ | K3: "\<lbrakk> evs3 \<in> kerbIV;
Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
Says Kas' A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
authTicket\<rbrace>) \<in> set evs3;
@@ -169,7 +166,7 @@
Theorems that exploit it have the suffix `_u', which stands for updated
protocol.
*)
- K4: "\<lbrakk> evs4 \<in> kerbIV; Key servK \<notin> used evs4; servK \<in> symKeys;
+ | K4: "\<lbrakk> evs4 \<in> kerbIV; Key servK \<notin> used evs4; servK \<in> symKeys;
B \<noteq> Tgs; authK \<in> symKeys;
Says A' Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
@@ -196,7 +193,7 @@
(*---------------------------------------------------------------------*)
(* FROM the initiator *)
- K5: "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
+ | K5: "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
Agent B\<rbrace>
@@ -213,7 +210,7 @@
(*---------------------------------------------------------------------*)
(* FROM the responder*)
- K6: "\<lbrakk> evs6 \<in> kerbIV;
+ | K6: "\<lbrakk> evs6 \<in> kerbIV;
Says A' B \<lbrace>
(Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
(Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
@@ -228,7 +225,7 @@
(*---------------------------------------------------------------------*)
(* Leaking an authK... *)
- Oops1: "\<lbrakk> evsO1 \<in> kerbIV; A \<noteq> Spy;
+ | Oops1: "\<lbrakk> evsO1 \<in> kerbIV; A \<noteq> Spy;
Says Kas A
(Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
authTicket\<rbrace>) \<in> set evsO1;
@@ -239,7 +236,7 @@
(*---------------------------------------------------------------------*)
(*Leaking a servK... *)
- Oops2: "\<lbrakk> evsO2 \<in> kerbIV; A \<noteq> Spy;
+ | Oops2: "\<lbrakk> evsO2 \<in> kerbIV; A \<noteq> Spy;
Says Tgs A
(Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
\<in> set evsO2;
--- a/src/HOL/Auth/KerberosIV_Gets.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy Wed Jul 11 11:14:51 2007 +0200
@@ -98,22 +98,19 @@
Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace> \<rbrace>)
\<in> set evs"
-consts
-
-kerbIV_gets :: "event list set"
-inductive "kerbIV_gets"
- intros
+inductive_set "kerbIV_gets" :: "event list set"
+ where
Nil: "[] \<in> kerbIV_gets"
- Fake: "\<lbrakk> evsf \<in> kerbIV_gets; X \<in> synth (analz (spies evsf)) \<rbrakk>
+ | Fake: "\<lbrakk> evsf \<in> kerbIV_gets; X \<in> synth (analz (spies evsf)) \<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> kerbIV_gets"
- Reception: "\<lbrakk> evsr \<in> kerbIV_gets; Says A B X \<in> set evsr \<rbrakk>
+ | Reception: "\<lbrakk> evsr \<in> kerbIV_gets; Says A B X \<in> set evsr \<rbrakk>
\<Longrightarrow> Gets B X # evsr \<in> kerbIV_gets"
(* FROM the initiator *)
- K1: "\<lbrakk> evs1 \<in> kerbIV_gets \<rbrakk>
+ | K1: "\<lbrakk> evs1 \<in> kerbIV_gets \<rbrakk>
\<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
\<in> kerbIV_gets"
@@ -124,7 +121,7 @@
(*---------------------------------------------------------------------*)
(*FROM Kas *)
- K2: "\<lbrakk> evs2 \<in> kerbIV_gets; Key authK \<notin> used evs2; authK \<in> symKeys;
+ | K2: "\<lbrakk> evs2 \<in> kerbIV_gets; Key authK \<notin> used evs2; authK \<in> symKeys;
Gets Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Kas A
(Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2),
@@ -140,7 +137,7 @@
(*---------------------------------------------------------------------*)
(* FROM the initiator *)
- K3: "\<lbrakk> evs3 \<in> kerbIV_gets;
+ | K3: "\<lbrakk> evs3 \<in> kerbIV_gets;
Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
Gets A (Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
authTicket\<rbrace>) \<in> set evs3;
@@ -160,7 +157,7 @@
Theorems that exploit it have the suffix `_u', which stands for updated
protocol.
*)
- K4: "\<lbrakk> evs4 \<in> kerbIV_gets; Key servK \<notin> used evs4; servK \<in> symKeys;
+ | K4: "\<lbrakk> evs4 \<in> kerbIV_gets; Key servK \<notin> used evs4; servK \<in> symKeys;
B \<noteq> Tgs; authK \<in> symKeys;
Gets Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
@@ -187,7 +184,7 @@
(*---------------------------------------------------------------------*)
(* FROM the initiator *)
- K5: "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
+ | K5: "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
Agent B\<rbrace>
@@ -204,7 +201,7 @@
(*---------------------------------------------------------------------*)
(* FROM the responder*)
- K6: "\<lbrakk> evs6 \<in> kerbIV_gets;
+ | K6: "\<lbrakk> evs6 \<in> kerbIV_gets;
Gets B \<lbrace>
(Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
(Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
@@ -219,7 +216,7 @@
(*---------------------------------------------------------------------*)
(* Leaking an authK... *)
- Oops1: "\<lbrakk> evsO1 \<in> kerbIV_gets; A \<noteq> Spy;
+ | Oops1: "\<lbrakk> evsO1 \<in> kerbIV_gets; A \<noteq> Spy;
Says Kas A
(Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta,
authTicket\<rbrace>) \<in> set evsO1;
@@ -230,7 +227,7 @@
(*---------------------------------------------------------------------*)
(*Leaking a servK... *)
- Oops2: "\<lbrakk> evsO2 \<in> kerbIV_gets; A \<noteq> Spy;
+ | Oops2: "\<lbrakk> evsO2 \<in> kerbIV_gets; A \<noteq> Spy;
Says Tgs A
(Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
\<in> set evsO2;
--- a/src/HOL/Auth/KerberosV.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/KerberosV.thy Wed Jul 11 11:14:51 2007 +0200
@@ -100,24 +100,21 @@
Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, tt\<rbrace> \<rbrace>
\<in> set evs"
-consts
-
-kerbV :: "event list set"
-inductive "kerbV"
- intros
+inductive_set kerbV :: "event list set"
+ where
Nil: "[] \<in> kerbV"
- Fake: "\<lbrakk> evsf \<in> kerbV; X \<in> synth (analz (spies evsf)) \<rbrakk>
+ | Fake: "\<lbrakk> evsf \<in> kerbV; X \<in> synth (analz (spies evsf)) \<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> kerbV"
(*Authentication phase*)
- KV1: "\<lbrakk> evs1 \<in> kerbV \<rbrakk>
+ | KV1: "\<lbrakk> evs1 \<in> kerbV \<rbrakk>
\<Longrightarrow> Says A Kas \<lbrace>Agent A, Agent Tgs, Number (CT evs1)\<rbrace> # evs1
\<in> kerbV"
(*Unlike version IV, authTicket is not re-encrypted*)
- KV2: "\<lbrakk> evs2 \<in> kerbV; Key authK \<notin> used evs2; authK \<in> symKeys;
+ | KV2: "\<lbrakk> evs2 \<in> kerbV; Key authK \<notin> used evs2; authK \<in> symKeys;
Says A' Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Kas A \<lbrace>
Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number (CT evs2)\<rbrace>,
@@ -126,7 +123,7 @@
(* Authorisation phase *)
- KV3: "\<lbrakk> evs3 \<in> kerbV; A \<noteq> Kas; A \<noteq> Tgs;
+ | KV3: "\<lbrakk> evs3 \<in> kerbV; A \<noteq> Kas; A \<noteq> Tgs;
Says A Kas \<lbrace>Agent A, Agent Tgs, Number T1\<rbrace> \<in> set evs3;
Says Kas' A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
authTicket\<rbrace> \<in> set evs3;
@@ -136,7 +133,7 @@
(Crypt authK \<lbrace>Agent A, Number (CT evs3)\<rbrace>),
Agent B\<rbrace> # evs3 \<in> kerbV"
(*Unlike version IV, servTicket is not re-encrypted*)
- KV4: "\<lbrakk> evs4 \<in> kerbV; Key servK \<notin> used evs4; servK \<in> symKeys;
+ | KV4: "\<lbrakk> evs4 \<in> kerbV; Key servK \<notin> used evs4; servK \<in> symKeys;
B \<noteq> Tgs; authK \<in> symKeys;
Says A' Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
@@ -154,7 +151,7 @@
(*Service phase*)
- KV5: "\<lbrakk> evs5 \<in> kerbV; authK \<in> symKeys; servK \<in> symKeys;
+ | KV5: "\<lbrakk> evs5 \<in> kerbV; authK \<in> symKeys; servK \<in> symKeys;
A \<noteq> Kas; A \<noteq> Tgs;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
@@ -168,7 +165,7 @@
Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
# evs5 \<in> kerbV"
- KV6: "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
+ | KV6: "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
Says A' B \<lbrace>
(Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>),
(Crypt servK \<lbrace>Agent A, Number T3\<rbrace>)\<rbrace>
@@ -182,7 +179,7 @@
(* Leaking an authK... *)
- Oops1:"\<lbrakk> evsO1 \<in> kerbV; A \<noteq> Spy;
+ | Oops1:"\<lbrakk> evsO1 \<in> kerbV; A \<noteq> Spy;
Says Kas A \<lbrace>Crypt (shrK A) \<lbrace>Key authK, Agent Tgs, Number Ta\<rbrace>,
authTicket\<rbrace> \<in> set evsO1;
expiredAK Ta evsO1 \<rbrakk>
@@ -190,7 +187,7 @@
# evsO1 \<in> kerbV"
(*Leaking a servK... *)
- Oops2: "\<lbrakk> evsO2 \<in> kerbV; A \<noteq> Spy;
+ | Oops2: "\<lbrakk> evsO2 \<in> kerbV; A \<noteq> Spy;
Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
servTicket\<rbrace> \<in> set evsO2;
expiredSK Ts evsO2 \<rbrakk>
--- a/src/HOL/Auth/Kerberos_BAN.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy Wed Jul 11 11:14:51 2007 +0200
@@ -71,22 +71,21 @@
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
-consts bankerberos :: "event list set"
-inductive "bankerberos"
- intros
+inductive_set bankerberos :: "event list set"
+ where
Nil: "[] \<in> bankerberos"
- Fake: "\<lbrakk> evsf \<in> bankerberos; X \<in> synth (analz (spies evsf)) \<rbrakk>
+ | Fake: "\<lbrakk> evsf \<in> bankerberos; X \<in> synth (analz (spies evsf)) \<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
- BK1: "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
+ | BK1: "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
\<in> bankerberos"
- BK2: "\<lbrakk> evs2 \<in> bankerberos; Key K \<notin> used evs2; K \<in> symKeys;
+ | BK2: "\<lbrakk> evs2 \<in> bankerberos; Key K \<notin> used evs2; K \<in> symKeys;
Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Server A
(Crypt (shrK A)
@@ -95,7 +94,7 @@
# evs2 \<in> bankerberos"
- BK3: "\<lbrakk> evs3 \<in> bankerberos;
+ | BK3: "\<lbrakk> evs3 \<in> bankerberos;
Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
\<in> set evs3;
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
@@ -104,7 +103,7 @@
# evs3 \<in> bankerberos"
- BK4: "\<lbrakk> evs4 \<in> bankerberos;
+ | BK4: "\<lbrakk> evs4 \<in> bankerberos;
Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
(Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
\<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
@@ -112,7 +111,7 @@
\<in> bankerberos"
(*Old session keys may become compromised*)
- Oops: "\<lbrakk> evso \<in> bankerberos;
+ | Oops: "\<lbrakk> evso \<in> bankerberos;
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
\<in> set evso;
expiredK Tk evso \<rbrakk>
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Wed Jul 11 11:14:51 2007 +0200
@@ -63,24 +63,23 @@
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
-consts bankerb_gets :: "event list set"
-inductive "bankerb_gets"
- intros
+inductive_set bankerb_gets :: "event list set"
+ where
Nil: "[] \<in> bankerb_gets"
- Fake: "\<lbrakk> evsf \<in> bankerb_gets; X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
+ | Fake: "\<lbrakk> evsf \<in> bankerb_gets; X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
- Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
+ | Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
\<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
- BK1: "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
+ | BK1: "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
\<in> bankerb_gets"
- BK2: "\<lbrakk> evs2 \<in> bankerb_gets; Key K \<notin> used evs2; K \<in> symKeys;
+ | BK2: "\<lbrakk> evs2 \<in> bankerb_gets; Key K \<notin> used evs2; K \<in> symKeys;
Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Server A
(Crypt (shrK A)
@@ -89,7 +88,7 @@
# evs2 \<in> bankerb_gets"
- BK3: "\<lbrakk> evs3 \<in> bankerb_gets;
+ | BK3: "\<lbrakk> evs3 \<in> bankerb_gets;
Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
\<in> set evs3;
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
@@ -98,7 +97,7 @@
# evs3 \<in> bankerb_gets"
- BK4: "\<lbrakk> evs4 \<in> bankerb_gets;
+ | BK4: "\<lbrakk> evs4 \<in> bankerb_gets;
Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
(Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
\<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
@@ -106,7 +105,7 @@
\<in> bankerb_gets"
(*Old session keys may become compromised*)
- Oops: "\<lbrakk> evso \<in> bankerb_gets;
+ | Oops: "\<lbrakk> evso \<in> bankerb_gets;
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
\<in> set evso;
expiredK Tk evso \<rbrakk>
--- a/src/HOL/Auth/Message.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Message.thy Wed Jul 11 11:14:51 2007 +0200
@@ -72,13 +72,14 @@
subsubsection{*Inductive Definition of All Parts" of a Message*}
-consts parts :: "msg set => msg set"
-inductive "parts H"
- intros
+inductive_set
+ parts :: "msg set => msg set"
+ for H :: "msg set"
+ where
Inj [intro]: "X \<in> H ==> X \<in> parts H"
- Fst: "{|X,Y|} \<in> parts H ==> X \<in> parts H"
- Snd: "{|X,Y|} \<in> parts H ==> Y \<in> parts H"
- Body: "Crypt K X \<in> parts H ==> X \<in> parts H"
+ | Fst: "{|X,Y|} \<in> parts H ==> X \<in> parts H"
+ | Snd: "{|X,Y|} \<in> parts H ==> Y \<in> parts H"
+ | Body: "Crypt K X \<in> parts H ==> X \<in> parts H"
text{*Monotonicity*}
@@ -335,13 +336,14 @@
messages, including keys. A form of downward closure. Pairs can
be taken apart; messages decrypted with known keys. *}
-consts analz :: "msg set => msg set"
-inductive "analz H"
- intros
+inductive_set
+ analz :: "msg set => msg set"
+ for H :: "msg set"
+ where
Inj [intro,simp] : "X \<in> H ==> X \<in> analz H"
- Fst: "{|X,Y|} \<in> analz H ==> X \<in> analz H"
- Snd: "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
- Decrypt [dest]:
+ | Fst: "{|X,Y|} \<in> analz H ==> X \<in> analz H"
+ | Snd: "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
+ | Decrypt [dest]:
"[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
@@ -460,14 +462,14 @@
analz (insert (Crypt K X) H) \<subseteq>
insert (Crypt K X) (analz (insert X H))"
apply (rule subsetI)
-apply (erule_tac xa = x in analz.induct, auto)
+apply (erule_tac x = x in analz.induct, auto)
done
lemma lemma2: "Key (invKey K) \<in> analz H ==>
insert (Crypt K X) (analz (insert X H)) \<subseteq>
analz (insert (Crypt K X) H)"
apply auto
-apply (erule_tac xa = x in analz.induct, auto)
+apply (erule_tac x = x in analz.induct, auto)
apply (blast intro: analz_insertI analz.Decrypt)
done
@@ -579,15 +581,16 @@
encrypted with known keys. Agent names are public domain.
Numbers can be guessed, but Nonces cannot be. *}
-consts synth :: "msg set => msg set"
-inductive "synth H"
- intros
+inductive_set
+ synth :: "msg set => msg set"
+ for H :: "msg set"
+ where
Inj [intro]: "X \<in> H ==> X \<in> synth H"
- Agent [intro]: "Agent agt \<in> synth H"
- Number [intro]: "Number n \<in> synth H"
- Hash [intro]: "X \<in> synth H ==> Hash X \<in> synth H"
- MPair [intro]: "[|X \<in> synth H; Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
- Crypt [intro]: "[|X \<in> synth H; Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
+ | Agent [intro]: "Agent agt \<in> synth H"
+ | Number [intro]: "Number n \<in> synth H"
+ | Hash [intro]: "X \<in> synth H ==> Hash X \<in> synth H"
+ | MPair [intro]: "[|X \<in> synth H; Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
+ | Crypt [intro]: "[|X \<in> synth H; Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
text{*Monotonicity*}
lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
--- a/src/HOL/Auth/NS_Public.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/NS_Public.thy Wed Jul 11 11:14:51 2007 +0200
@@ -11,32 +11,30 @@
theory NS_Public imports Public begin
-consts ns_public :: "event list set"
-
-inductive ns_public
- intros
+inductive_set ns_public :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> ns_public"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>
+ | Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> ns_public"
(*Alice initiates a protocol run, sending a nonce to Bob*)
- NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>
+ | NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>
\<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
# evs1 \<in> ns_public"
(*Bob responds to Alice's message with a further nonce*)
- NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;
+ | NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;
Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
\<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
# evs2 \<in> ns_public"
(*Alice proves her existence by sending NB back to Bob.*)
- NS3: "\<lbrakk>evs3 \<in> ns_public;
+ | NS3: "\<lbrakk>evs3 \<in> ns_public;
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>)
\<in> set evs3\<rbrakk>
--- a/src/HOL/Auth/NS_Public_Bad.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy Wed Jul 11 11:14:51 2007 +0200
@@ -15,32 +15,30 @@
theory NS_Public_Bad imports Public begin
-consts ns_public :: "event list set"
-
-inductive ns_public
- intros
+inductive_set ns_public :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> ns_public"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>
+ | Fake: "\<lbrakk>evsf \<in> ns_public; X \<in> synth (analz (spies evsf))\<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> ns_public"
(*Alice initiates a protocol run, sending a nonce to Bob*)
- NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>
+ | NS1: "\<lbrakk>evs1 \<in> ns_public; Nonce NA \<notin> used evs1\<rbrakk>
\<Longrightarrow> Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>)
# evs1 \<in> ns_public"
(*Bob responds to Alice's message with a further nonce*)
- NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;
+ | NS2: "\<lbrakk>evs2 \<in> ns_public; Nonce NB \<notin> used evs2;
Says A' B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs2\<rbrakk>
\<Longrightarrow> Says B A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>)
# evs2 \<in> ns_public"
(*Alice proves her existence by sending NB back to Bob.*)
- NS3: "\<lbrakk>evs3 \<in> ns_public;
+ | NS3: "\<lbrakk>evs3 \<in> ns_public;
Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs3;
Says B' A (Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs3\<rbrakk>
\<Longrightarrow> Says A B (Crypt (pubEK B) (Nonce NB)) # evs3 \<in> ns_public"
--- a/src/HOL/Auth/NS_Shared.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/NS_Shared.thy Wed Jul 11 11:14:51 2007 +0200
@@ -25,26 +25,25 @@
X \<notin> parts (spies (takeWhile (% z. z \<noteq> Says A B Y) (rev evs)))"
-consts ns_shared :: "event list set"
-inductive "ns_shared"
- intros
+inductive_set ns_shared :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> ns_shared"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk>
+| Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk>
\<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
(*Alice initiates a protocol run, requesting to talk to any B*)
- NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk>
+| NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk>
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
(*Server's response to Alice's message.
!! It may respond more than once to A's request !!
Server doesn't know who the true sender is, hence the A' in
the sender field.*)
- NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys;
+| NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys;
Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
\<Longrightarrow> Says Server A
(Crypt (shrK A)
@@ -54,14 +53,14 @@
(*We can't assume S=Server. Agent A "remembers" her nonce.
Need A \<noteq> Server because we allow messages to self.*)
- NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server;
+| NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server;
Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
\<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
(*Bob's nonce exchange. He does not know who the message came
from, but responds to A because she is mentioned inside.*)
- NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys;
+| NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys;
Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
\<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
@@ -70,7 +69,7 @@
We do NOT send NB-1 or similar as the Spy cannot spoof such things.
Letting the Spy add or subtract 1 lets him send all nonces.
Instead we distinguish the messages by sending the nonce twice.*)
- NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys;
+| NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys;
Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
\<in> set evs5\<rbrakk>
@@ -79,7 +78,7 @@
(*This message models possible leaks of session keys.
The two Nonces identify the protocol run: the rule insists upon
the true senders in order to make them accurate.*)
- Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso;
+| Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso;
Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
\<in> set evso\<rbrakk>
\<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
--- a/src/HOL/Auth/OtwayRees.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/OtwayRees.thy Wed Jul 11 11:14:51 2007 +0200
@@ -14,31 +14,30 @@
This is the original version, which encrypts Nonce NB.*}
-consts otway :: "event list set"
-inductive "otway"
- intros
+inductive_set otway :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> otway"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |]
+ | Fake: "[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> otway"
(*A message that has been sent can be received by the
intended recipient.*)
- Reception: "[| evsr \<in> otway; Says A B X \<in>set evsr |]
+ | Reception: "[| evsr \<in> otway; Says A B X \<in>set evsr |]
==> Gets B X # evsr \<in> otway"
(*Alice initiates a protocol run*)
- OR1: "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
+ | OR1: "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
==> Says A B {|Nonce NA, Agent A, Agent B,
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
# evs1 : otway"
(*Bob's response to Alice's message. Note that NB is encrypted.*)
- OR2: "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
+ | OR2: "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
Gets B {|Nonce NA, Agent A, Agent B, X|} : set evs2 |]
==> Says B Server
{|Nonce NA, Agent A, Agent B, X,
@@ -49,7 +48,7 @@
(*The Server receives Bob's message and checks that the three NAs
match. Then he sends a new session key to Bob with a packet for
forwarding to Alice.*)
- OR3: "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
+ | OR3: "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
Gets Server
{|Nonce NA, Agent A, Agent B,
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
@@ -64,7 +63,7 @@
(*Bob receives the Server's (?) message and compares the Nonces with
those in the message he previously sent the Server.
Need B \<noteq> Server because we allow messages to self.*)
- OR4: "[| evs4 \<in> otway; B \<noteq> Server;
+ | OR4: "[| evs4 \<in> otway; B \<noteq> Server;
Says B Server {|Nonce NA, Agent A, Agent B, X',
Crypt (shrK B)
{|Nonce NA, Nonce NB, Agent A, Agent B|}|}
@@ -75,7 +74,7 @@
(*This message models possible leaks of session keys. The nonces
identify the protocol run.*)
- Oops: "[| evso \<in> otway;
+ | Oops: "[| evso \<in> otway;
Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
: set evso |]
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso : otway"
--- a/src/HOL/Auth/OtwayReesBella.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy Wed Jul 11 11:14:51 2007 +0200
@@ -16,31 +16,30 @@
updated protocol makes no use of the session key to encrypt but informs A that
B knows it.*}
-consts orb :: "event list set"
-inductive "orb"
- intros
+inductive_set orb :: "event list set"
+ where
Nil: "[]\<in> orb"
- Fake: "\<lbrakk>evsa\<in> orb; X\<in> synth (analz (knows Spy evsa))\<rbrakk>
+| Fake: "\<lbrakk>evsa\<in> orb; X\<in> synth (analz (knows Spy evsa))\<rbrakk>
\<Longrightarrow> Says Spy B X # evsa \<in> orb"
- Reception: "\<lbrakk>evsr\<in> orb; Says A B X \<in> set evsr\<rbrakk>
+| Reception: "\<lbrakk>evsr\<in> orb; Says A B X \<in> set evsr\<rbrakk>
\<Longrightarrow> Gets B X # evsr \<in> orb"
- OR1: "\<lbrakk>evs1\<in> orb; Nonce NA \<notin> used evs1\<rbrakk>
+| OR1: "\<lbrakk>evs1\<in> orb; Nonce NA \<notin> used evs1\<rbrakk>
\<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,
Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
# evs1 \<in> orb"
- OR2: "\<lbrakk>evs2\<in> orb; Nonce NB \<notin> used evs2;
+| OR2: "\<lbrakk>evs2\<in> orb; Nonce NB \<notin> used evs2;
Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
\<Longrightarrow> Says B Server
\<lbrace>Nonce M, Agent A, Agent B, X,
Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
# evs2 \<in> orb"
- OR3: "\<lbrakk>evs3\<in> orb; Key KAB \<notin> used evs3;
+| OR3: "\<lbrakk>evs3\<in> orb; Key KAB \<notin> used evs3;
Gets Server
\<lbrace>Nonce M, Agent A, Agent B,
Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,
@@ -53,7 +52,7 @@
(*B can only check that the message he is bouncing is a ciphertext*)
(*Sending M back is omitted*)
- OR4: "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>;
+| OR4: "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>;
Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',
Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
\<in> set evs4;
@@ -62,7 +61,7 @@
\<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
- Oops: "\<lbrakk>evso\<in> orb;
+| Oops: "\<lbrakk>evso\<in> orb;
Says Server B \<lbrace>Nonce M,
Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
Nonce NB, Key KAB\<rbrace>\<rbrace>
--- a/src/HOL/Auth/OtwayRees_AN.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/OtwayRees_AN.thy Wed Jul 11 11:14:51 2007 +0200
@@ -22,34 +22,33 @@
IEEE Trans. SE 22 (1)
*}
-consts otway :: "event list set"
-inductive "otway"
- intros
+inductive_set otway :: "event list set"
+ where
Nil: --{*The empty trace*}
"[] \<in> otway"
- Fake: --{*The Spy may say anything he can say. The sender field is correct,
+ | Fake: --{*The Spy may say anything he can say. The sender field is correct,
but agents don't use that information.*}
"[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> otway"
- Reception: --{*A message that has been sent can be received by the
+ | Reception: --{*A message that has been sent can be received by the
intended recipient.*}
"[| evsr \<in> otway; Says A B X \<in>set evsr |]
==> Gets B X # evsr \<in> otway"
- OR1: --{*Alice initiates a protocol run*}
+ | OR1: --{*Alice initiates a protocol run*}
"evs1 \<in> otway
==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
- OR2: --{*Bob's response to Alice's message.*}
+ | OR2: --{*Bob's response to Alice's message.*}
"[| evs2 \<in> otway;
Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
# evs2 \<in> otway"
- OR3: --{*The Server receives Bob's message. Then he sends a new
+ | OR3: --{*The Server receives Bob's message. Then he sends a new
session key to Bob with a packet for forwarding to Alice.*}
"[| evs3 \<in> otway; Key KAB \<notin> used evs3;
Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
@@ -59,7 +58,7 @@
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key KAB|}|}
# evs3 \<in> otway"
- OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
+ | OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
those in the message he previously sent the Server.
Need @{term "B \<noteq> Server"} because we allow messages to self.*}
"[| evs4 \<in> otway; B \<noteq> Server;
@@ -68,7 +67,7 @@
\<in>set evs4 |]
==> Says B A X # evs4 \<in> otway"
- Oops: --{*This message models possible leaks of session keys. The nonces
+ | Oops: --{*This message models possible leaks of session keys. The nonces
identify the protocol run.*}
"[| evso \<in> otway;
Says Server B
--- a/src/HOL/Auth/OtwayRees_Bad.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.thy Wed Jul 11 11:14:51 2007 +0200
@@ -19,30 +19,29 @@
the protocol is open to a middleperson attack. Attempting to prove some key
lemmas indicates the possibility of this attack.*}
-consts otway :: "event list set"
-inductive "otway"
- intros
+inductive_set otway :: "event list set"
+ where
Nil: --{*The empty trace*}
"[] \<in> otway"
- Fake: --{*The Spy may say anything he can say. The sender field is correct,
+ | Fake: --{*The Spy may say anything he can say. The sender field is correct,
but agents don't use that information.*}
"[| evsf \<in> otway; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> otway"
- Reception: --{*A message that has been sent can be received by the
+ | Reception: --{*A message that has been sent can be received by the
intended recipient.*}
"[| evsr \<in> otway; Says A B X \<in>set evsr |]
==> Gets B X # evsr \<in> otway"
- OR1: --{*Alice initiates a protocol run*}
+ | OR1: --{*Alice initiates a protocol run*}
"[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
==> Says A B {|Nonce NA, Agent A, Agent B,
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
# evs1 \<in> otway"
- OR2: --{*Bob's response to Alice's message.
+ | OR2: --{*Bob's response to Alice's message.
This variant of the protocol does NOT encrypt NB.*}
"[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
@@ -51,7 +50,7 @@
Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
# evs2 \<in> otway"
- OR3: --{*The Server receives Bob's message and checks that the three NAs
+ | OR3: --{*The Server receives Bob's message and checks that the three NAs
match. Then he sends a new session key to Bob with a packet for
forwarding to Alice.*}
"[| evs3 \<in> otway; Key KAB \<notin> used evs3;
@@ -67,7 +66,7 @@
Crypt (shrK B) {|Nonce NB, Key KAB|}|}
# evs3 \<in> otway"
- OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
+ | OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
those in the message he previously sent the Server.
Need @{term "B \<noteq> Server"} because we allow messages to self.*}
"[| evs4 \<in> otway; B \<noteq> Server;
@@ -78,7 +77,7 @@
\<in> set evs4 |]
==> Says B A {|Nonce NA, X|} # evs4 \<in> otway"
- Oops: --{*This message models possible leaks of session keys. The nonces
+ | Oops: --{*This message models possible leaks of session keys. The nonces
identify the protocol run.*}
"[| evso \<in> otway;
Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
--- a/src/HOL/Auth/Recur.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Recur.thy Wed Jul 11 11:14:51 2007 +0200
@@ -17,16 +17,17 @@
who receives one.
Perhaps the two session keys could be bundled into a single message.
*)
-consts respond :: "event list => (msg*msg*key)set"
-inductive "respond evs" (*Server's response to the nested message*)
- intros
+inductive_set (*Server's response to the nested message*)
+ respond :: "event list => (msg*msg*key)set"
+ for evs :: "event list"
+ where
One: "Key KAB \<notin> used evs
==> (Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, END|},
{|Crypt (shrK A) {|Key KAB, Agent B, Nonce NA|}, END|},
KAB) \<in> respond evs"
(*The most recent session key is passed up to the caller*)
- Cons: "[| (PA, RA, KAB) \<in> respond evs;
+ | Cons: "[| (PA, RA, KAB) \<in> respond evs;
Key KBC \<notin> used evs; Key KBC \<notin> parts {RA};
PA = Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, P|} |]
==> (Hash[Key(shrK B)] {|Agent B, Agent C, Nonce NB, PA|},
@@ -40,50 +41,50 @@
(*Induction over "respond" can be difficult due to the complexity of the
subgoals. Set "responses" captures the general form of certificates.
*)
-consts responses :: "event list => msg set"
-inductive "responses evs"
- intros
+inductive_set
+ responses :: "event list => msg set"
+ for evs :: "event list"
+ where
(*Server terminates lists*)
Nil: "END \<in> responses evs"
- Cons: "[| RA \<in> responses evs; Key KAB \<notin> used evs |]
+ | Cons: "[| RA \<in> responses evs; Key KAB \<notin> used evs |]
==> {|Crypt (shrK B) {|Key KAB, Agent A, Nonce NB|},
RA|} \<in> responses evs"
-consts recur :: "event list set"
-inductive "recur"
- intros
+inductive_set recur :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> recur"
(*The spy MAY say anything he CAN say. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> recur; X \<in> synth (analz (knows Spy evsf)) |]
+ | Fake: "[| evsf \<in> recur; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> recur"
(*Alice initiates a protocol run.
END is a placeholder to terminate the nesting.*)
- RA1: "[| evs1 \<in> recur; Nonce NA \<notin> used evs1 |]
+ | RA1: "[| evs1 \<in> recur; Nonce NA \<notin> used evs1 |]
==> Says A B (Hash[Key(shrK A)] {|Agent A, Agent B, Nonce NA, END|})
# evs1 \<in> recur"
(*Bob's response to Alice's message. C might be the Server.
We omit PA = {|XA, Agent A, Agent B, Nonce NA, P|} because
it complicates proofs, so B may respond to any message at all!*)
- RA2: "[| evs2 \<in> recur; Nonce NB \<notin> used evs2;
+ | RA2: "[| evs2 \<in> recur; Nonce NB \<notin> used evs2;
Says A' B PA \<in> set evs2 |]
==> Says B C (Hash[Key(shrK B)] {|Agent B, Agent C, Nonce NB, PA|})
# evs2 \<in> recur"
(*The Server receives Bob's message and prepares a response.*)
- RA3: "[| evs3 \<in> recur; Says B' Server PB \<in> set evs3;
+ | RA3: "[| evs3 \<in> recur; Says B' Server PB \<in> set evs3;
(PB,RB,K) \<in> respond evs3 |]
==> Says Server B RB # evs3 \<in> recur"
(*Bob receives the returned message and compares the Nonces with
those in the message he previously sent the Server.*)
- RA4: "[| evs4 \<in> recur;
+ | RA4: "[| evs4 \<in> recur;
Says B C {|XH, Agent B, Agent C, Nonce NB,
XA, Agent A, Agent B, Nonce NA, P|} \<in> set evs4;
Says C' B {|Crypt (shrK B) {|Key KBC, Agent C, Nonce NB|},
@@ -350,7 +351,7 @@
lemma respond_certificate:
"(Hash[Key(shrK A)] {|Agent A, B, NA, P|}, RA, K) \<in> respond evs
==> Crypt (shrK A) {|Key K, B, NA|} \<in> parts {RA}"
-apply (ind_cases "(X, RA, K) \<in> respond evs")
+apply (ind_cases "(Hash[Key (shrK A)] \<lbrace>Agent A, B, NA, P\<rbrace>, RA, K) \<in> respond evs")
apply simp_all
done
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Jul 11 11:14:51 2007 +0200
@@ -36,15 +36,14 @@
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
-consts sr :: "event list set"
-inductive "sr"
- intros
+inductive_set sr :: "event list set"
+ where
Nil: "[]\<in> sr"
- Fake: "\<lbrakk> evsF\<in> sr; X\<in> synth (analz (knows Spy evsF));
+ | Fake: "\<lbrakk> evsF\<in> sr; X\<in> synth (analz (knows Spy evsF));
illegalUse(Card B) \<rbrakk>
\<Longrightarrow> Says Spy A X #
Inputs Spy (Card B) X # evsF \<in> sr"
@@ -52,24 +51,24 @@
(*In general this rule causes the assumption Card B \<notin> cloned
in most guarantees for B - starting with confidentiality -
otherwise pairK_confidential could not apply*)
- Forge:
+ | Forge:
"\<lbrakk> evsFo \<in> sr; Nonce Nb \<in> analz (knows Spy evsFo);
Key (pairK(A,B)) \<in> knows Spy evsFo \<rbrakk>
\<Longrightarrow> Notes Spy (Key (sesK(Nb,pairK(A,B)))) # evsFo \<in> sr"
- Reception: "\<lbrakk> evsR\<in> sr; Says A B X \<in> set evsR \<rbrakk>
+ | Reception: "\<lbrakk> evsR\<in> sr; Says A B X \<in> set evsR \<rbrakk>
\<Longrightarrow> Gets B X # evsR \<in> sr"
(*A AND THE SERVER *)
- SR1: "\<lbrakk> evs1\<in> sr; A \<noteq> Server\<rbrakk>
+ | SR1: "\<lbrakk> evs1\<in> sr; A \<noteq> Server\<rbrakk>
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace>
# evs1 \<in> sr"
- SR2: "\<lbrakk> evs2\<in> sr;
+ | SR2: "\<lbrakk> evs2\<in> sr;
Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Server A \<lbrace>Nonce (Pairkey(A,B)),
Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>
@@ -82,7 +81,7 @@
(*A AND HER CARD*)
(*A cannot decrypt the verifier for she dosn't know shrK A,
but the pairkey is recognisable*)
- SR3: "\<lbrakk> evs3\<in> sr; legalUse(Card A);
+ | SR3: "\<lbrakk> evs3\<in> sr; legalUse(Card A);
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
Gets A \<lbrace>Nonce Pk, Certificate\<rbrace> \<in> set evs3 \<rbrakk>
\<Longrightarrow> Inputs A (Card A) (Agent A)
@@ -93,7 +92,7 @@
the server*)
(*The card outputs the nonce Na to A*)
- SR4: "\<lbrakk> evs4\<in> sr; A \<noteq> Server;
+ | SR4: "\<lbrakk> evs4\<in> sr; A \<noteq> Server;
Nonce Na \<notin> used evs4; legalUse(Card A);
Inputs A (Card A) (Agent A) \<in> set evs4 \<rbrakk>
\<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
@@ -101,9 +100,9 @@
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
- SR4Fake: "\<lbrakk> evs4F\<in> sr; Nonce Na \<notin> used evs4F;
- illegalUse(Card A);
- Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk>
+ | SR4Fake: "\<lbrakk> evs4F\<in> sr; Nonce Na \<notin> used evs4F;
+ illegalUse(Card A);
+ Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk>
\<Longrightarrow> Outpts (Card A) Spy \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
# evs4F \<in> sr"
@@ -111,7 +110,7 @@
(*A TOWARDS B*)
- SR5: "\<lbrakk> evs5\<in> sr;
+ | SR5: "\<lbrakk> evs5\<in> sr;
Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs5;
\<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
\<Longrightarrow> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> # evs5 \<in> sr"
@@ -122,13 +121,13 @@
(*B AND HIS CARD*)
- SR6: "\<lbrakk> evs6\<in> sr; legalUse(Card B);
+ | SR6: "\<lbrakk> evs6\<in> sr; legalUse(Card B);
Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs6 \<rbrakk>
\<Longrightarrow> Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace>
# evs6 \<in> sr"
(*B gets back from the card the session key and various verifiers*)
- SR7: "\<lbrakk> evs7\<in> sr;
+ | SR7: "\<lbrakk> evs7\<in> sr;
Nonce Nb \<notin> used evs7; legalUse(Card B); B \<noteq> Server;
K = sesK(Nb,pairK(A,B));
Key K \<notin> used evs7;
@@ -140,11 +139,11 @@
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
- SR7Fake: "\<lbrakk> evs7F\<in> sr; Nonce Nb \<notin> used evs7F;
- illegalUse(Card B);
- K = sesK(Nb,pairK(A,B));
- Key K \<notin> used evs7F;
- Inputs Spy (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7F \<rbrakk>
+ | SR7Fake: "\<lbrakk> evs7F\<in> sr; Nonce Nb \<notin> used evs7F;
+ illegalUse(Card B);
+ K = sesK(Nb,pairK(A,B));
+ Key K \<notin> used evs7F;
+ Inputs Spy (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs7F \<rbrakk>
\<Longrightarrow> Outpts (Card B) Spy \<lbrace>Nonce Nb, Key K,
Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,
Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
@@ -156,7 +155,7 @@
(*B TOWARDS A*)
(*having sent an input that mentions A is the only memory B relies on,
since the output doesn't mention A - lack of explicitness*)
- SR8: "\<lbrakk> evs8\<in> sr;
+ | SR8: "\<lbrakk> evs8\<in> sr;
Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs8;
Outpts (Card B) B \<lbrace>Nonce Nb, Key K,
Cert1, Cert2\<rbrace> \<in> set evs8 \<rbrakk>
@@ -168,7 +167,7 @@
(*A AND HER CARD*)
(*A cannot check the form of the verifiers - although I can prove the form of
Cert2 - and just feeds her card with what she's got*)
- SR9: "\<lbrakk> evs9\<in> sr; legalUse(Card A);
+ | SR9: "\<lbrakk> evs9\<in> sr; legalUse(Card A);
Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs9;
Outpts (Card A) A \<lbrace>Nonce Na, Cert2\<rbrace> \<in> set evs9;
Gets A \<lbrace>Nonce Nb, Cert3\<rbrace> \<in> set evs9;
@@ -179,7 +178,7 @@
# evs9 \<in> sr"
(*But the card will only give outputs to the inputs of the correct form*)
- SR10: "\<lbrakk> evs10\<in> sr; legalUse(Card A); A \<noteq> Server;
+ | SR10: "\<lbrakk> evs10\<in> sr; legalUse(Card A); A \<noteq> Server;
K = sesK(Nb,pairK(A,B));
Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb,
Nonce (Pairkey(A,B)),
@@ -193,16 +192,16 @@
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
-SR10Fake: "\<lbrakk> evs10F\<in> sr;
- illegalUse(Card A);
- K = sesK(Nb,pairK(A,B));
- Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb,
- Nonce (Pairkey(A,B)),
- Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)),
- Agent B\<rbrace>,
- Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,
- Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
- \<in> set evs10F \<rbrakk>
+ | SR10Fake: "\<lbrakk> evs10F\<in> sr;
+ illegalUse(Card A);
+ K = sesK(Nb,pairK(A,B));
+ Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb,
+ Nonce (Pairkey(A,B)),
+ Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)),
+ Agent B\<rbrace>,
+ Crypt (pairK(A,B)) \<lbrace>Nonce Na, Nonce Nb\<rbrace>,
+ Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
+ \<in> set evs10F \<rbrakk>
\<Longrightarrow> Outpts (Card A) Spy \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
# evs10F \<in> sr"
@@ -212,7 +211,7 @@
(*A TOWARDS B*)
(*having initiated with B is the only memory A relies on,
since the output doesn't mention B - lack of explicitness*)
- SR11: "\<lbrakk> evs11\<in> sr;
+ | SR11: "\<lbrakk> evs11\<in> sr;
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs11;
Outpts (Card A) A \<lbrace>Key K, Certificate\<rbrace> \<in> set evs11 \<rbrakk>
\<Longrightarrow> Says A B (Certificate)
@@ -222,13 +221,13 @@
(*Both peers may leak by accident the session keys obtained from their
cards*)
- Oops1:
+ | Oops1:
"\<lbrakk> evsO1 \<in> sr;
Outpts (Card B) B \<lbrace>Nonce Nb, Key K, Certificate,
Crypt (pairK(A,B)) (Nonce Nb)\<rbrace> \<in> set evsO1 \<rbrakk>
\<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO1 \<in> sr"
- Oops2:
+ | Oops2:
"\<lbrakk> evsO2 \<in> sr;
Outpts (Card A) A \<lbrace>Key K, Crypt (pairK(A,B)) (Nonce Nb)\<rbrace>
\<in> set evsO2 \<rbrakk>
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Jul 11 11:14:51 2007 +0200
@@ -42,15 +42,14 @@
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
-consts srb :: "event list set"
-inductive "srb"
- intros
+inductive_set srb :: "event list set"
+ where
Nil: "[]\<in> srb"
- Fake: "\<lbrakk> evsF \<in> srb; X \<in> synth (analz (knows Spy evsF));
+ | Fake: "\<lbrakk> evsF \<in> srb; X \<in> synth (analz (knows Spy evsF));
illegalUse(Card B) \<rbrakk>
\<Longrightarrow> Says Spy A X #
Inputs Spy (Card B) X # evsF \<in> srb"
@@ -58,24 +57,24 @@
(*In general this rule causes the assumption Card B \<notin> cloned
in most guarantees for B - starting with confidentiality -
otherwise pairK_confidential could not apply*)
- Forge:
+ | Forge:
"\<lbrakk> evsFo \<in> srb; Nonce Nb \<in> analz (knows Spy evsFo);
Key (pairK(A,B)) \<in> knows Spy evsFo \<rbrakk>
\<Longrightarrow> Notes Spy (Key (sesK(Nb,pairK(A,B)))) # evsFo \<in> srb"
- Reception: "\<lbrakk> evsrb\<in> srb; Says A B X \<in> set evsrb \<rbrakk>
+ | Reception: "\<lbrakk> evsrb\<in> srb; Says A B X \<in> set evsrb \<rbrakk>
\<Longrightarrow> Gets B X # evsrb \<in> srb"
(*A AND THE SERVER*)
- SR_U1: "\<lbrakk> evs1 \<in> srb; A \<noteq> Server \<rbrakk>
+ | SR_U1: "\<lbrakk> evs1 \<in> srb; A \<noteq> Server \<rbrakk>
\<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace>
# evs1 \<in> srb"
- SR_U2: "\<lbrakk> evs2 \<in> srb;
+ | SR_U2: "\<lbrakk> evs2 \<in> srb;
Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
\<Longrightarrow> Says Server A \<lbrace>Nonce (Pairkey(A,B)),
Crypt (shrK A) \<lbrace>Nonce (Pairkey(A,B)), Agent B\<rbrace>
@@ -88,7 +87,7 @@
(*A AND HER CARD*)
(*A cannot decrypt the verifier for she dosn't know shrK A,
but the pairkey is recognisable*)
- SR_U3: "\<lbrakk> evs3 \<in> srb; legalUse(Card A);
+ | SR_U3: "\<lbrakk> evs3 \<in> srb; legalUse(Card A);
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
Gets A \<lbrace>Nonce Pk, Certificate\<rbrace> \<in> set evs3 \<rbrakk>
\<Longrightarrow> Inputs A (Card A) (Agent A)
@@ -99,7 +98,7 @@
the server*)
(*The card outputs the nonce Na to A*)
- SR_U4: "\<lbrakk> evs4 \<in> srb;
+ | SR_U4: "\<lbrakk> evs4 \<in> srb;
Nonce Na \<notin> used evs4; legalUse(Card A); A \<noteq> Server;
Inputs A (Card A) (Agent A) \<in> set evs4 \<rbrakk>
\<Longrightarrow> Outpts (Card A) A \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
@@ -107,7 +106,7 @@
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
- SR_U4Fake: "\<lbrakk> evs4F \<in> srb; Nonce Na \<notin> used evs4F;
+ | SR_U4Fake: "\<lbrakk> evs4F \<in> srb; Nonce Na \<notin> used evs4F;
illegalUse(Card A);
Inputs Spy (Card A) (Agent A) \<in> set evs4F \<rbrakk>
\<Longrightarrow> Outpts (Card A) Spy \<lbrace>Nonce Na, Crypt (crdK (Card A)) (Nonce Na)\<rbrace>
@@ -117,7 +116,7 @@
(*A TOWARDS B*)
- SR_U5: "\<lbrakk> evs5 \<in> srb;
+ | SR_U5: "\<lbrakk> evs5 \<in> srb;
Outpts (Card A) A \<lbrace>Nonce Na, Certificate\<rbrace> \<in> set evs5;
\<forall> p q. Certificate \<noteq> \<lbrace>p, q\<rbrace> \<rbrakk>
\<Longrightarrow> Says A B \<lbrace>Agent A, Nonce Na\<rbrace> # evs5 \<in> srb"
@@ -128,12 +127,12 @@
(*B AND HIS CARD*)
- SR_U6: "\<lbrakk> evs6 \<in> srb; legalUse(Card B);
+ | SR_U6: "\<lbrakk> evs6 \<in> srb; legalUse(Card B);
Gets B \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs6 \<rbrakk>
\<Longrightarrow> Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace>
# evs6 \<in> srb"
(*B gets back from the card the session key and various verifiers*)
- SR_U7: "\<lbrakk> evs7 \<in> srb;
+ | SR_U7: "\<lbrakk> evs7 \<in> srb;
Nonce Nb \<notin> used evs7; legalUse(Card B); B \<noteq> Server;
K = sesK(Nb,pairK(A,B));
Key K \<notin> used evs7;
@@ -144,7 +143,7 @@
# evs7 \<in> srb"
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
- SR_U7Fake: "\<lbrakk> evs7F \<in> srb; Nonce Nb \<notin> used evs7F;
+ | SR_U7Fake: "\<lbrakk> evs7F \<in> srb; Nonce Nb \<notin> used evs7F;
illegalUse(Card B);
K = sesK(Nb,pairK(A,B));
Key K \<notin> used evs7F;
@@ -160,7 +159,7 @@
(*B TOWARDS A*)
(*having sent an input that mentions A is the only memory B relies on,
since the output doesn't mention A - lack of explicitness*)
- SR_U8: "\<lbrakk> evs8 \<in> srb;
+ | SR_U8: "\<lbrakk> evs8 \<in> srb;
Inputs B (Card B) \<lbrace>Agent A, Nonce Na\<rbrace> \<in> set evs8;
Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K,
Cert1, Cert2\<rbrace> \<in> set evs8 \<rbrakk>
@@ -172,7 +171,7 @@
(*A AND HER CARD*)
(*A cannot check the form of the verifiers - although I can prove the form of
Cert2 - and just feeds her card with what she's got*)
- SR_U9: "\<lbrakk> evs9 \<in> srb; legalUse(Card A);
+ | SR_U9: "\<lbrakk> evs9 \<in> srb; legalUse(Card A);
Gets A \<lbrace>Nonce Pk, Cert1\<rbrace> \<in> set evs9;
Outpts (Card A) A \<lbrace>Nonce Na, Cert2\<rbrace> \<in> set evs9;
Gets A \<lbrace>Nonce Nb, Cert3\<rbrace> \<in> set evs9;
@@ -182,7 +181,7 @@
Cert1, Cert3, Cert2\<rbrace>
# evs9 \<in> srb"
(*But the card will only give outputs to the inputs of the correct form*)
- SR_U10: "\<lbrakk> evs10 \<in> srb; legalUse(Card A); A \<noteq> Server;
+ | SR_U10: "\<lbrakk> evs10 \<in> srb; legalUse(Card A); A \<noteq> Server;
K = sesK(Nb,pairK(A,B));
Inputs A (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb,
Nonce (Pairkey(A,B)),
@@ -196,7 +195,7 @@
# evs10 \<in> srb"
(*The card can be exploited by the spy*)
(*because of the assumptions on the card, A is certainly not server nor spy*)
-SR_U10Fake: "\<lbrakk> evs10F \<in> srb;
+ | SR_U10Fake: "\<lbrakk> evs10F \<in> srb;
illegalUse(Card A);
K = sesK(Nb,pairK(A,B));
Inputs Spy (Card A) \<lbrace>Agent B, Nonce Na, Nonce Nb,
@@ -216,7 +215,7 @@
(*A TOWARDS B*)
(*having initiated with B is the only memory A relies on,
since the output doesn't mention B - lack of explicitness*)
- SR_U11: "\<lbrakk> evs11 \<in> srb;
+ | SR_U11: "\<lbrakk> evs11 \<in> srb;
Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs11;
Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace>
\<in> set evs11 \<rbrakk>
@@ -227,13 +226,13 @@
(*Both peers may leak by accident the session keys obtained from their
cards*)
- Oops1:
+ | Oops1:
"\<lbrakk> evsO1 \<in> srb;
Outpts (Card B) B \<lbrace>Nonce Nb, Agent A, Key K, Cert1, Cert2\<rbrace>
\<in> set evsO1 \<rbrakk>
\<Longrightarrow> Notes Spy \<lbrace>Key K, Nonce Nb, Agent A, Agent B\<rbrace> # evsO1 \<in> srb"
- Oops2:
+ | Oops2:
"\<lbrakk> evsO2 \<in> srb;
Outpts (Card A) A \<lbrace>Agent B, Nonce Nb, Key K, Certificate\<rbrace>
\<in> set evsO2 \<rbrakk>
--- a/src/HOL/Auth/TLS.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/TLS.thy Wed Jul 11 11:14:51 2007 +0200
@@ -99,25 +99,24 @@
sessionK_neq_shrK [iff]: "sessionK nonces \<noteq> shrK A"
-consts tls :: "event list set"
-inductive tls
- intros
+inductive_set tls :: "event list set"
+ where
Nil: --{*The initial, empty trace*}
"[] \<in> tls"
- Fake: --{*The Spy may say anything he can say. The sender field is correct,
+ | Fake: --{*The Spy may say anything he can say. The sender field is correct,
but agents don't use that information.*}
"[| evsf \<in> tls; X \<in> synth (analz (spies evsf)) |]
==> Says Spy B X # evsf \<in> tls"
- SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
+ | SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
to available nonces*}
"[| evsSK \<in> tls;
{Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
==> Notes Spy {| Nonce (PRF(M,NA,NB)),
Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
- ClientHello:
+ | ClientHello:
--{*(7.4.1.2)
PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
It is uninterpreted but will be confirmed in the FINISHED messages.
@@ -129,7 +128,7 @@
==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
# evsCH \<in> tls"
- ServerHello:
+ | ServerHello:
--{*7.4.1.3 of the TLS Internet-Draft
PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
SERVER CERTIFICATE (7.4.2) is always present.
@@ -139,11 +138,11 @@
\<in> set evsSH |]
==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH \<in> tls"
- Certificate:
+ | Certificate:
--{*SERVER (7.4.2) or CLIENT (7.4.6) CERTIFICATE.*}
"evsC \<in> tls ==> Says B A (certificate B (pubK B)) # evsC \<in> tls"
- ClientKeyExch:
+ | ClientKeyExch:
--{*CLIENT KEY EXCHANGE (7.4.7).
The client, A, chooses PMS, the PREMASTER SECRET.
She encrypts PMS using the supplied KB, which ought to be pubK B.
@@ -158,7 +157,7 @@
# Notes A {|Agent B, Nonce PMS|}
# evsCX \<in> tls"
- CertVerify:
+ | CertVerify:
--{*The optional Certificate Verify (7.4.8) message contains the
specific components listed in the security analysis, F.1.1.2.
It adds the pre-master-secret, which is also essential!
@@ -174,7 +173,7 @@
among other things. The master-secret is PRF(PMS,NA,NB).
Either party may send its message first.*}
- ClientFinished:
+ | ClientFinished:
--{*The occurrence of Notes A {|Agent B, Nonce PMS|} stops the
rule's applying when the Spy has satisfied the "Says A B" by
repaying messages sent by the true client; in that case, the
@@ -193,7 +192,7 @@
Nonce NB, Number PB, Agent B|}))
# evsCF \<in> tls"
- ServerFinished:
+ | ServerFinished:
--{*Keeping A' and A'' distinct means B cannot even check that the
two messages originate from the same source. *}
"[| evsSF \<in> tls;
@@ -208,7 +207,7 @@
Nonce NB, Number PB, Agent B|}))
# evsSF \<in> tls"
- ClientAccepts:
+ | ClientAccepts:
--{*Having transmitted ClientFinished and received an identical
message encrypted with serverK, the client stores the parameters
needed to resume this session. The "Notes A ..." premise is
@@ -224,7 +223,7 @@
==>
Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA \<in> tls"
- ServerAccepts:
+ | ServerAccepts:
--{*Having transmitted ServerFinished and received an identical
message encrypted with clientK, the server stores the parameters
needed to resume this session. The "Says A'' B ..." premise is
@@ -241,7 +240,7 @@
==>
Notes B {|Number SID, Agent A, Agent B, Nonce M|} # evsSA \<in> tls"
- ClientResume:
+ | ClientResume:
--{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
message using the new nonces and stored MASTER SECRET.*}
"[| evsCR \<in> tls;
@@ -254,7 +253,7 @@
Nonce NB, Number PB, Agent B|}))
# evsCR \<in> tls"
- ServerResume:
+ | ServerResume:
--{*Resumption (7.3): If B finds the @{text SESSION_ID} then he can
send a FINISHED message using the recovered MASTER SECRET*}
"[| evsSR \<in> tls;
@@ -267,7 +266,7 @@
Nonce NB, Number PB, Agent B|})) # evsSR
\<in> tls"
- Oops:
+ | Oops:
--{*The most plausible compromise is of an old session key. Losing
the MASTER SECRET or PREMASTER SECRET is more serious but
rather unlikely. The assumption @{term "A\<noteq>Spy"} is essential:
--- a/src/HOL/Auth/WooLam.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/WooLam.thy Wed Jul 11 11:14:51 2007 +0200
@@ -18,9 +18,8 @@
Computer Security Foundations Workshop
*}
-consts woolam :: "event list set"
-inductive woolam
- intros
+inductive_set woolam :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> woolam"
@@ -29,20 +28,20 @@
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> woolam; X \<in> synth (analz (spies evsf)) |]
+ | Fake: "[| evsf \<in> woolam; X \<in> synth (analz (spies evsf)) |]
==> Says Spy B X # evsf \<in> woolam"
(*Alice initiates a protocol run*)
- WL1: "evs1 \<in> woolam ==> Says A B (Agent A) # evs1 \<in> woolam"
+ | WL1: "evs1 \<in> woolam ==> Says A B (Agent A) # evs1 \<in> woolam"
(*Bob responds to Alice's message with a challenge.*)
- WL2: "[| evs2 \<in> woolam; Says A' B (Agent A) \<in> set evs2 |]
+ | WL2: "[| evs2 \<in> woolam; Says A' B (Agent A) \<in> set evs2 |]
==> Says B A (Nonce NB) # evs2 \<in> woolam"
(*Alice responds to Bob's challenge by encrypting NB with her key.
B is *not* properly determined -- Alice essentially broadcasts
her reply.*)
- WL3: "[| evs3 \<in> woolam;
+ | WL3: "[| evs3 \<in> woolam;
Says A B (Agent A) \<in> set evs3;
Says B' A (Nonce NB) \<in> set evs3 |]
==> Says A B (Crypt (shrK A) (Nonce NB)) # evs3 \<in> woolam"
@@ -51,13 +50,13 @@
the messages are shown in chronological order, for clarity.
But here, exchanging the two events would cause the lemma
WL4_analz_spies to pick up the wrong assumption!*)
- WL4: "[| evs4 \<in> woolam;
+ | WL4: "[| evs4 \<in> woolam;
Says A' B X \<in> set evs4;
Says A'' B (Agent A) \<in> set evs4 |]
==> Says B Server {|Agent A, Agent B, X|} # evs4 \<in> woolam"
(*Server decrypts Alice's response for Bob.*)
- WL5: "[| evs5 \<in> woolam;
+ | WL5: "[| evs5 \<in> woolam;
Says B' Server {|Agent A, Agent B, Crypt (shrK A) (Nonce NB)|}
\<in> set evs5 |]
==> Says Server B (Crypt (shrK B) {|Agent A, Nonce NB|})
--- a/src/HOL/Auth/Yahalom.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Yahalom.thy Wed Jul 11 11:14:51 2007 +0200
@@ -15,29 +15,28 @@
This theory has the prototypical example of a secrecy relation, KeyCryptNonce.
*}
-consts yahalom :: "event list set"
-inductive "yahalom"
- intros
+inductive_set yahalom :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> yahalom"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
+ | Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> yahalom"
(*A message that has been sent can be received by the
intended recipient.*)
- Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
+ | Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> yahalom"
(*Alice initiates a protocol run*)
- YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
+ | YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
==> Says A B {|Agent A, Nonce NA|} # evs1 \<in> yahalom"
(*Bob's response to Alice's message.*)
- YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
+ | YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
Gets B {|Agent A, Nonce NA|} \<in> set evs2 |]
==> Says B Server
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
@@ -45,7 +44,7 @@
(*The Server receives Bob's message. He responds by sending a
new session key to Alice, with a packet for forwarding to Bob.*)
- YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys;
+ | YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys;
Gets Server
{|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
\<in> set evs3 |]
@@ -54,7 +53,7 @@
Crypt (shrK B) {|Agent A, Key KAB|}|}
# evs3 \<in> yahalom"
- YM4:
+ | YM4:
--{*Alice receives the Server's (?) message, checks her Nonce, and
uses the new session key to send Bob his Nonce. The premise
@{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
@@ -68,7 +67,7 @@
(*This message models possible leaks of session keys. The Nonces
identify the protocol run. Quoting Server here ensures they are
correct.*)
- Oops: "[| evso \<in> yahalom;
+ | Oops: "[| evso \<in> yahalom;
Says Server A {|Crypt (shrK A)
{|Agent B, Key K, Nonce NA, Nonce NB|},
X|} \<in> set evso |]
--- a/src/HOL/Auth/Yahalom2.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Yahalom2.thy Wed Jul 11 11:14:51 2007 +0200
@@ -19,29 +19,28 @@
This theory has the prototypical example of a secrecy relation, KeyCryptNonce.
*}
-consts yahalom :: "event list set"
-inductive "yahalom"
- intros
+inductive_set yahalom :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> yahalom"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
+ | Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> yahalom"
(*A message that has been sent can be received by the
intended recipient.*)
- Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
+ | Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> yahalom"
(*Alice initiates a protocol run*)
- YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
+ | YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
==> Says A B {|Agent A, Nonce NA|} # evs1 \<in> yahalom"
(*Bob's response to Alice's message.*)
- YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
+ | YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
Gets B {|Agent A, Nonce NA|} \<in> set evs2 |]
==> Says B Server
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
@@ -50,7 +49,7 @@
(*The Server receives Bob's message. He responds by sending a
new session key to Alice, with a certificate for forwarding to Bob.
Both agents are quoted in the 2nd certificate to prevent attacks!*)
- YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3;
+ | YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3;
Gets Server {|Agent B, Nonce NB,
Crypt (shrK B) {|Agent A, Nonce NA|}|}
\<in> set evs3 |]
@@ -62,7 +61,7 @@
(*Alice receives the Server's (?) message, checks her Nonce, and
uses the new session key to send Bob his Nonce.*)
- YM4: "[| evs4 \<in> yahalom;
+ | YM4: "[| evs4 \<in> yahalom;
Gets A {|Nonce NB, Crypt (shrK A) {|Agent B, Key K, Nonce NA|},
X|} \<in> set evs4;
Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
@@ -71,7 +70,7 @@
(*This message models possible leaks of session keys. The nonces
identify the protocol run. Quoting Server here ensures they are
correct. *)
- Oops: "[| evso \<in> yahalom;
+ | Oops: "[| evso \<in> yahalom;
Says Server A {|Nonce NB,
Crypt (shrK A) {|Agent B, Key K, Nonce NA|},
X|} \<in> set evso |]
--- a/src/HOL/Auth/Yahalom_Bad.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/Yahalom_Bad.thy Wed Jul 11 11:14:51 2007 +0200
@@ -14,29 +14,28 @@
The issues are discussed in lcp's LICS 2000 invited lecture.
*}
-consts yahalom :: "event list set"
-inductive "yahalom"
- intros
+inductive_set yahalom :: "event list set"
+ where
(*Initial trace is empty*)
Nil: "[] \<in> yahalom"
(*The spy MAY say anything he CAN say. We do not expect him to
invent new nonces here, but he can also use NS1. Common to
all similar protocols.*)
- Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
+ | Fake: "[| evsf \<in> yahalom; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> yahalom"
(*A message that has been sent can be received by the
intended recipient.*)
- Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
+ | Reception: "[| evsr \<in> yahalom; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> yahalom"
(*Alice initiates a protocol run*)
- YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
+ | YM1: "[| evs1 \<in> yahalom; Nonce NA \<notin> used evs1 |]
==> Says A B {|Agent A, Nonce NA|} # evs1 \<in> yahalom"
(*Bob's response to Alice's message.*)
- YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
+ | YM2: "[| evs2 \<in> yahalom; Nonce NB \<notin> used evs2;
Gets B {|Agent A, Nonce NA|} \<in> set evs2 |]
==> Says B Server
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
@@ -44,7 +43,7 @@
(*The Server receives Bob's message. He responds by sending a
new session key to Alice, with a packet for forwarding to Bob.*)
- YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys;
+ | YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3; KAB \<in> symKeys;
Gets Server
{|Agent B, Nonce NB, Crypt (shrK B) {|Agent A, Nonce NA|}|}
\<in> set evs3 |]
@@ -56,7 +55,7 @@
(*Alice receives the Server's (?) message, checks her Nonce, and
uses the new session key to send Bob his Nonce. The premise
A \<noteq> Server is needed to prove Says_Server_not_range.*)
- YM4: "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
+ | YM4: "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
\<in> set evs4;
Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
--- a/src/HOL/Auth/ZhouGollmann.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Auth/ZhouGollmann.thy Wed Jul 11 11:14:51 2007 +0200
@@ -29,29 +29,27 @@
declare broken_def [simp]
-consts zg :: "event list set"
-
-inductive zg
- intros
+inductive_set zg :: "event list set"
+ where
Nil: "[] \<in> zg"
- Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |]
+| Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |]
==> Says Spy B X # evsf \<in> zg"
-Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
+| Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
(*L is fresh for honest agents.
We don't require K to be fresh because we don't bother to prove secrecy!
We just assume that the protocol's objective is to deliver K fairly,
rather than to keep M secret.*)
- ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m);
+| ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m);
K \<in> symKeys;
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
(*B must check that NRO is A's signature to learn the sender's name*)
- ZG2: "[| evs2 \<in> zg;
+| ZG2: "[| evs2 \<in> zg;
Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
@@ -59,7 +57,7 @@
(*A must check that NRR is B's signature to learn the sender's name;
without spy, the matching label would be enough*)
- ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
+| ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
@@ -73,7 +71,7 @@
give con_K to the Spy. This makes the threat model more dangerous, while
also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
@{term "K \<noteq> priK TTP"}. *)
- ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
+| ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
\<in> set evs4;
sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
--- a/src/HOL/HoareParallel/OG_Hoare.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/HoareParallel/OG_Hoare.thy Wed Jul 11 11:14:51 2007 +0200
@@ -36,46 +36,40 @@
"interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow>
interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
-consts ann_hoare :: "('a ann_com \<times> 'a assn) set"
-syntax "_ann_hoare" :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(2\<turnstile> _// _)" [60,90] 45)
-translations "\<turnstile> c q" \<rightleftharpoons> "(c, q) \<in> ann_hoare"
-
-consts oghoare :: "('a assn \<times> 'a com \<times> 'a assn) set"
-syntax "_oghoare" :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>- _//_//_)" [90,55,90] 50)
-translations "\<parallel>- p c q" \<rightleftharpoons> "(p, c, q) \<in> oghoare"
-
-inductive oghoare ann_hoare
-intros
+inductive
+ oghoare :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>- _//_//_)" [90,55,90] 50)
+ and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(2\<turnstile> _// _)" [60,90] 45)
+where
AnnBasic: "r \<subseteq> {s. f s \<in> q} \<Longrightarrow> \<turnstile> (AnnBasic r f) q"
- AnnSeq: "\<lbrakk> \<turnstile> c0 pre c1; \<turnstile> c1 q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnSeq c0 c1) q"
+| AnnSeq: "\<lbrakk> \<turnstile> c0 pre c1; \<turnstile> c1 q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnSeq c0 c1) q"
- AnnCond1: "\<lbrakk> r \<inter> b \<subseteq> pre c1; \<turnstile> c1 q; r \<inter> -b \<subseteq> pre c2; \<turnstile> c2 q\<rbrakk>
+| AnnCond1: "\<lbrakk> r \<inter> b \<subseteq> pre c1; \<turnstile> c1 q; r \<inter> -b \<subseteq> pre c2; \<turnstile> c2 q\<rbrakk>
\<Longrightarrow> \<turnstile> (AnnCond1 r b c1 c2) q"
- AnnCond2: "\<lbrakk> r \<inter> b \<subseteq> pre c; \<turnstile> c q; r \<inter> -b \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnCond2 r b c) q"
+| AnnCond2: "\<lbrakk> r \<inter> b \<subseteq> pre c; \<turnstile> c q; r \<inter> -b \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnCond2 r b c) q"
- AnnWhile: "\<lbrakk> r \<subseteq> i; i \<inter> b \<subseteq> pre c; \<turnstile> c i; i \<inter> -b \<subseteq> q \<rbrakk>
+| AnnWhile: "\<lbrakk> r \<subseteq> i; i \<inter> b \<subseteq> pre c; \<turnstile> c i; i \<inter> -b \<subseteq> q \<rbrakk>
\<Longrightarrow> \<turnstile> (AnnWhile r b i c) q"
- AnnAwait: "\<lbrakk> atom_com c; \<parallel>- (r \<inter> b) c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r b c) q"
+| AnnAwait: "\<lbrakk> atom_com c; \<parallel>- (r \<inter> b) c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r b c) q"
- AnnConseq: "\<lbrakk>\<turnstile> c q; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<turnstile> c q'"
+| AnnConseq: "\<lbrakk>\<turnstile> c q; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<turnstile> c q'"
- Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk>
+| Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk>
\<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i))))
Parallel Ts
(\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
- Basic: "\<parallel>- {s. f s \<in>q} (Basic f) q"
+| Basic: "\<parallel>- {s. f s \<in>q} (Basic f) q"
- Seq: "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q "
+| Seq: "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q "
- Cond: "\<lbrakk> \<parallel>- (p \<inter> b) c1 q; \<parallel>- (p \<inter> -b) c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Cond b c1 c2) q"
+| Cond: "\<lbrakk> \<parallel>- (p \<inter> b) c1 q; \<parallel>- (p \<inter> -b) c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Cond b c1 c2) q"
- While: "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)"
+| While: "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)"
- Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- p' c q'"
+| Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- p' c q'"
section {* Soundness *}
(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE
@@ -147,13 +141,13 @@
subsection {* Soundness of the System for Component Programs *}
inductive_cases ann_transition_cases:
- "(None,s) -1\<rightarrow> t"
- "(Some (AnnBasic r f),s) -1\<rightarrow> t"
- "(Some (AnnSeq c1 c2), s) -1\<rightarrow> t"
- "(Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> t"
- "(Some (AnnCond2 r b c), s) -1\<rightarrow> t"
- "(Some (AnnWhile r b I c), s) -1\<rightarrow> t"
- "(Some (AnnAwait r b c),s) -1\<rightarrow> t"
+ "(None,s) -1\<rightarrow> (c', s')"
+ "(Some (AnnBasic r f),s) -1\<rightarrow> (c', s')"
+ "(Some (AnnSeq c1 c2), s) -1\<rightarrow> (c', s')"
+ "(Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (c', s')"
+ "(Some (AnnCond2 r b c), s) -1\<rightarrow> (c', s')"
+ "(Some (AnnWhile r b I c), s) -1\<rightarrow> (c', s')"
+ "(Some (AnnAwait r b c),s) -1\<rightarrow> (c', s')"
text {* Strong Soundness for Component Programs:*}
@@ -174,7 +168,7 @@
apply(clarify,simp,clarify,rule_tac x=qa in exI,fast)
done
-lemma Help: "(transition \<inter> {(v,v,u). True}) = (transition)"
+lemma Help: "(transition \<inter> {(x,y). True}) = (transition)"
apply force
done
@@ -412,7 +406,7 @@
apply clarify
apply(drule Parallel_length_post_PStar)
apply clarify
-apply (ind_cases "(Parallel Ts, s) -P1\<rightarrow> (Parallel Rs, t)")
+apply (ind_cases "(Parallel Ts, s) -P1\<rightarrow> (Parallel Rs, t)" for Ts s Rs t)
apply(rule conjI)
apply clarify
apply(case_tac "i=j")
--- a/src/HOL/HoareParallel/OG_Tran.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/HoareParallel/OG_Tran.thy Wed Jul 11 11:14:51 2007 +0200
@@ -19,71 +19,72 @@
subsection {* The Transition Relation *}
-consts
+inductive_set
ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"
- transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
-
-syntax
- "_ann_transition" :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
- ("_ -1\<rightarrow> _"[81,81] 100)
- "_ann_transition_n" :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
- \<Rightarrow> bool" ("_ -_\<rightarrow> _"[81,81] 100)
- "_ann_transition_*" :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
- ("_ -*\<rightarrow> _"[81,81] 100)
+ and transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
+ and ann_transition' :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
+ ("_ -1\<rightarrow> _"[81,81] 100)
+ and transition' :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
+ ("_ -P1\<rightarrow> _"[81,81] 100)
+ and transitions :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
+ ("_ -P*\<rightarrow> _"[81,81] 100)
+where
+ "con_0 -1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition"
+| "con_0 -P1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition"
+| "con_0 -P*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition\<^sup>*"
+
+| AnnBasic: "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
+
+| AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow>
+ (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"
+| AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow>
+ (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"
+
+| AnnCond1T: "s \<in> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"
+| AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"
- "_transition" :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool" ("_ -P1\<rightarrow> _"[81,81] 100)
- "_transition_n" :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
- ("_ -P_\<rightarrow> _"[81,81,81] 100)
- "_transition_*" :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool" ("_ -P*\<rightarrow> _"[81,81] 100)
+| AnnCond2T: "s \<in> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"
+| AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"
+
+| AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"
+| AnnWhileT: "s \<in> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow>
+ (Some (AnnSeq c (AnnWhile i b i c)), s)"
+
+| AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
+ (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)"
+
+| Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
+ \<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
+
+| Basic: "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"
+
+| Seq1: "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"
+| Seq2: "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"
+
+| CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"
+| CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"
+
+| WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"
+| WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"
+
+monos "rtrancl_mono"
text {* The corresponding syntax translations are: *}
-translations
- "con_0 -1\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition"
- "con_0 -n\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition^n"
- "con_0 -*\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> ann_transition\<^sup>*"
-
- "con_0 -P1\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition"
- "con_0 -Pn\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition^n"
- "con_0 -P*\<rightarrow> con_1" \<rightleftharpoons> "(con_0, con_1) \<in> transition\<^sup>*"
-
-inductive ann_transition transition
-intros
- AnnBasic: "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
-
- AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow>
- (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"
- AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow>
- (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"
-
- AnnCond1T: "s \<in> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"
- AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"
+abbreviation
+ ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
+ \<Rightarrow> bool" ("_ -_\<rightarrow> _"[81,81] 100) where
+ "con_0 -n\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition^n"
- AnnCond2T: "s \<in> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"
- AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"
-
- AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"
- AnnWhileT: "s \<in> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow>
- (Some (AnnSeq c (AnnWhile i b i c)), s)"
-
- AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
- (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)"
-
- Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
- \<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
+abbreviation
+ ann_transitions :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
+ ("_ -*\<rightarrow> _"[81,81] 100) where
+ "con_0 -*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition\<^sup>*"
- Basic: "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"
-
- Seq1: "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"
- Seq2: "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"
-
- CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"
- CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"
-
- WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"
- WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"
-
-monos "rtrancl_mono"
+abbreviation
+ transition_n :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
+ ("_ -P_\<rightarrow> _"[81,81,81] 100) where
+ "con_0 -Pn\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition^n"
subsection {* Definition of Semantics *}
--- a/src/HOL/HoareParallel/RG_Hoare.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/HoareParallel/RG_Hoare.thy Wed Jul 11 11:14:51 2007 +0200
@@ -11,36 +11,31 @@
stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"
"stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)"
-consts rghoare :: "('a rgformula) set"
-syntax
- "_rghoare" :: "['a com, 'a set, ('a \<times> 'a) set, ('a \<times> 'a) set, 'a set] \<Rightarrow> bool"
- ("\<turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
-translations
- "\<turnstile> P sat [pre, rely, guar, post]" \<rightleftharpoons> "(P, pre, rely, guar, post) \<in> rghoare"
-
-inductive rghoare
-intros
+inductive
+ rghoare :: "['a com, 'a set, ('a \<times> 'a) set, ('a \<times> 'a) set, 'a set] \<Rightarrow> bool"
+ ("\<turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+where
Basic: "\<lbrakk> pre \<subseteq> {s. f s \<in> post}; {(s,t). s \<in> pre \<and> (t=f s \<or> t=s)} \<subseteq> guar;
stable pre rely; stable post rely \<rbrakk>
\<Longrightarrow> \<turnstile> Basic f sat [pre, rely, guar, post]"
- Seq: "\<lbrakk> \<turnstile> P sat [pre, rely, guar, mid]; \<turnstile> Q sat [mid, rely, guar, post] \<rbrakk>
+| Seq: "\<lbrakk> \<turnstile> P sat [pre, rely, guar, mid]; \<turnstile> Q sat [mid, rely, guar, post] \<rbrakk>
\<Longrightarrow> \<turnstile> Seq P Q sat [pre, rely, guar, post]"
- Cond: "\<lbrakk> stable pre rely; \<turnstile> P1 sat [pre \<inter> b, rely, guar, post];
+| Cond: "\<lbrakk> stable pre rely; \<turnstile> P1 sat [pre \<inter> b, rely, guar, post];
\<turnstile> P2 sat [pre \<inter> -b, rely, guar, post]; \<forall>s. (s,s)\<in>guar \<rbrakk>
\<Longrightarrow> \<turnstile> Cond b P1 P2 sat [pre, rely, guar, post]"
- While: "\<lbrakk> stable pre rely; (pre \<inter> -b) \<subseteq> post; stable post rely;
+| While: "\<lbrakk> stable pre rely; (pre \<inter> -b) \<subseteq> post; stable post rely;
\<turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s,s)\<in>guar \<rbrakk>
\<Longrightarrow> \<turnstile> While b P sat [pre, rely, guar, post]"
- Await: "\<lbrakk> stable pre rely; stable post rely;
+| Await: "\<lbrakk> stable pre rely; stable post rely;
\<forall>V. \<turnstile> P sat [pre \<inter> b \<inter> {V}, {(s, t). s = t},
UNIV, {s. (V, s) \<in> guar} \<inter> post] \<rbrakk>
\<Longrightarrow> \<turnstile> Await b P sat [pre, rely, guar, post]"
- Conseq: "\<lbrakk> pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post;
+| Conseq: "\<lbrakk> pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post;
\<turnstile> P sat [pre', rely', guar', post'] \<rbrakk>
\<Longrightarrow> \<turnstile> P sat [pre, rely, guar, post]"
@@ -60,14 +55,10 @@
types 'a par_rgformula = "('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
-consts par_rghoare :: "('a par_rgformula) set"
-syntax
- "_par_rghoare" :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" ("\<turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
-translations
- "\<turnstile> Ps SAT [pre, rely, guar, post]" \<rightleftharpoons> "(Ps, pre, rely, guar, post) \<in> par_rghoare"
-
-inductive par_rghoare
-intros
+inductive
+ par_rghoare :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
+ ("\<turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+where
Parallel:
"\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j<length xs \<and> j\<noteq>i}. Guar(xs!j)) \<subseteq> Rely(xs!i);
(\<Union>j\<in>{j. j<length xs}. Guar(xs!j)) \<subseteq> guar;
@@ -113,22 +104,22 @@
lemma takecptn_is_cptn [rule_format, elim!]:
"\<forall>j. c \<in> cptn \<longrightarrow> take (Suc j) c \<in> cptn"
apply(induct "c")
- apply(force elim: cptn.elims)
+ apply(force elim: cptn.cases)
apply clarify
apply(case_tac j)
apply simp
apply(rule CptnOne)
apply simp
-apply(force intro:cptn.intros elim:cptn.elims)
+apply(force intro:cptn.intros elim:cptn.cases)
done
lemma dropcptn_is_cptn [rule_format,elim!]:
"\<forall>j<length c. c \<in> cptn \<longrightarrow> drop j c \<in> cptn"
apply(induct "c")
- apply(force elim: cptn.elims)
+ apply(force elim: cptn.cases)
apply clarify
apply(case_tac j,simp+)
-apply(erule cptn.elims)
+apply(erule cptn.cases)
apply simp
apply force
apply force
@@ -137,20 +128,20 @@
lemma takepar_cptn_is_par_cptn [rule_format,elim]:
"\<forall>j. c \<in> par_cptn \<longrightarrow> take (Suc j) c \<in> par_cptn"
apply(induct "c")
- apply(force elim: cptn.elims)
+ apply(force elim: cptn.cases)
apply clarify
apply(case_tac j,simp)
apply(rule ParCptnOne)
-apply(force intro:par_cptn.intros elim:par_cptn.elims)
+apply(force intro:par_cptn.intros elim:par_cptn.cases)
done
lemma droppar_cptn_is_par_cptn [rule_format]:
"\<forall>j<length c. c \<in> par_cptn \<longrightarrow> drop j c \<in> par_cptn"
apply(induct "c")
- apply(force elim: par_cptn.elims)
+ apply(force elim: par_cptn.cases)
apply clarify
apply(case_tac j,simp+)
-apply(erule par_cptn.elims)
+apply(erule par_cptn.cases)
apply simp
apply force
apply force
@@ -165,16 +156,16 @@
"\<forall>s. (None, s)#xs \<in> cptn \<longrightarrow> (\<forall>i<length xs. ((None, s)#xs)!i -e\<rightarrow> xs!i)"
apply(induct xs,simp+)
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply simp
apply(case_tac i,simp)
apply(rule Env)
apply simp
-apply(force elim:ctran.elims)
+apply(force elim:ctran.cases)
done
lemma cptn_not_empty [simp]:"[] \<notin> cptn"
-apply(force elim:cptn.elims)
+apply(force elim:cptn.cases)
done
lemma etran_or_ctran [rule_format]:
@@ -183,7 +174,7 @@
\<longrightarrow> x!i -e\<rightarrow> x!Suc i"
apply(induct x,simp)
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp)
apply(rule Env)
apply simp
@@ -202,10 +193,10 @@
apply(induct x)
apply simp
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp+)
apply(case_tac i,simp)
- apply(force elim:etran.elims)
+ apply(force elim:etran.cases)
apply simp
done
@@ -221,7 +212,7 @@
"\<lbrakk> (None, s) # xs \<in>cptn; i<length xs\<rbrakk> \<Longrightarrow> \<not> ((None, s) # xs) ! i -c\<rightarrow> xs ! i"
apply(frule not_ctran_None,simp)
apply(case_tac i,simp)
- apply(force elim:etran.elims)
+ apply(force elim:etranE)
apply simp
apply(rule etran_or_ctran2_disjI2,simp_all)
apply(force intro:tl_of_cptn_is_cptn)
@@ -241,9 +232,9 @@
(\<forall>i. j\<le>i \<and> i<k \<longrightarrow> x!i -e\<rightarrow> x!Suc i) \<longrightarrow> snd(x!k)\<in>p \<and> fst(x!j)=fst(x!k)"
apply(induct x)
apply clarify
- apply(force elim:cptn.elims)
+ apply(force elim:cptn.cases)
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply simp
apply(case_tac k,simp,simp)
apply(case_tac j,simp)
@@ -274,7 +265,7 @@
apply(case_tac k,simp,simp)
apply(case_tac j)
apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
- apply(erule etran.elims,simp)
+ apply(erule etran.cases,simp)
apply(erule_tac x="nata" in allE)
apply(erule_tac x="nat" and P="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((Q, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((Q, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
@@ -295,7 +286,7 @@
apply(induct x,simp)
apply simp
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp+)
apply clarify
apply(case_tac j,simp)
@@ -305,12 +296,12 @@
apply simp
apply(case_tac i)
apply(case_tac j,simp,simp)
- apply(erule ctran.elims,simp_all)
+ apply(erule ctran.cases,simp_all)
apply(force elim: not_ctran_None)
-apply(ind_cases "((Some (Basic f), sa), Q, t) \<in> ctran")
+apply(ind_cases "((Some (Basic f), sa), Q, t) \<in> ctran" for sa Q t)
apply simp
apply(drule_tac i=nat in not_ctran_None,simp)
-apply(erule etran.elims,simp)
+apply(erule etranE,simp)
done
lemma exists_ctran_Basic_None [rule_format]:
@@ -319,7 +310,7 @@
apply(induct x,simp)
apply simp
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp,simp)
apply(erule_tac x=nat in allE,simp)
apply clarify
@@ -349,7 +340,7 @@
apply clarify
apply(drule_tac s="Some (Basic f)" in sym,simp)
apply(thin_tac "\<forall>j. ?H j")
- apply(force elim:ctran.elims)
+ apply(force elim:ctran.cases)
apply clarify
apply(simp add:cp_def)
apply clarify
@@ -368,7 +359,7 @@
apply simp
apply(drule_tac s="Some (Basic f)" in sym,simp)
apply(case_tac "x!Suc j",simp)
-apply(rule ctran.elims,simp)
+apply(rule ctran.cases,simp)
apply(simp_all)
apply(drule_tac c=sa in subsetD,simp)
apply clarify
@@ -389,7 +380,7 @@
(\<forall>j. Suc j<length x \<longrightarrow> i\<noteq>j \<longrightarrow> x!j -e\<rightarrow> x!Suc j)"
apply(induct x,simp+)
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp+)
apply clarify
apply(case_tac j,simp)
@@ -399,11 +390,11 @@
apply simp
apply(case_tac i)
apply(case_tac j,simp,simp)
- apply(erule ctran.elims,simp_all)
+ apply(erule ctran.cases,simp_all)
apply(force elim: not_ctran_None)
-apply(ind_cases "((Some (Await b c), sa), Q, t) \<in> ctran",simp)
+apply(ind_cases "((Some (Await b c), sa), Q, t) \<in> ctran" for sa Q t,simp)
apply(drule_tac i=nat in not_ctran_None,simp)
-apply(erule etran.elims,simp)
+apply(erule etranE,simp)
done
lemma exists_ctran_Await_None [rule_format]:
@@ -411,7 +402,7 @@
\<longrightarrow> i<length x \<longrightarrow> fst(x!i)=None \<longrightarrow> (\<exists>j<i. x!j -c\<rightarrow> x!Suc j)"
apply(induct x,simp+)
apply clarify
-apply(erule cptn.elims,simp)
+apply(erule cptn.cases,simp)
apply(case_tac i,simp+)
apply(erule_tac x=nat in allE,simp)
apply clarify
@@ -440,7 +431,7 @@
apply force
apply(simp add:cp_def)
apply(case_tac l)
- apply(force elim:cptn.elims)
+ apply(force elim:cptn.cases)
apply simp
apply(erule CptnComp)
apply clarify
@@ -466,7 +457,7 @@
apply(erule_tac i=i in unique_ctran_Await,force,simp_all)
apply(simp add:cp_def)
--{* here starts the different part. *}
- apply(erule ctran.elims,simp_all)
+ apply(erule ctran.cases,simp_all)
apply(drule Star_imp_cptn)
apply clarify
apply(erule_tac x=sa in allE)
@@ -476,7 +467,7 @@
apply (simp add:cp_def)
apply clarify
apply(erule_tac x=ia and P="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply simp
apply clarify
apply(simp add:cp_def)
@@ -496,7 +487,7 @@
apply simp
apply(drule_tac s="Some (Await b P)" in sym,simp)
apply(case_tac "x!Suc j",simp)
-apply(rule ctran.elims,simp)
+apply(rule ctran.cases,simp)
apply(simp_all)
apply(drule Star_imp_cptn)
apply clarify
@@ -507,7 +498,7 @@
apply (simp add:cp_def)
apply clarify
apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply simp
apply clarify
apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all)
@@ -544,7 +535,7 @@
apply (simp add:assum_def)
apply(frule_tac j=0 and k="m" and p=pre in stability,simp+)
apply(erule_tac m="Suc m" in etran_or_ctran,simp+)
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
apply(erule_tac x="sa" in allE)
apply(drule_tac c="drop (Suc m) x" in subsetD)
apply simp
@@ -616,7 +607,7 @@
apply(simp (no_asm_use) add:lift_def)
apply clarify
apply(erule_tac x="Suc i" in allE, simp)
- apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \<in> ctran")
+ apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \<in> ctran" for Pa sa t)
apply(rule_tac x="(Some P, sa) # xs" in exI, simp add:cptn_iff_cptn_mod lift_def)
apply(erule_tac x="length xs" in allE, simp)
apply(simp only:Cons_lift_append)
@@ -649,7 +640,7 @@
apply(rule conjI,erule CptnEnv)
apply(simp (no_asm_use) add:lift_def)
apply(rule_tac x=ys in exI,simp)
-apply(ind_cases "((Some (Seq Pa Q), sa), t) \<in> ctran")
+apply(ind_cases "((Some (Seq Pa Q), sa), t) \<in> ctran" for Pa sa t)
apply simp
apply(rule_tac x="(Some Pa, sa)#[(None, ta)]" in exI,simp)
apply(rule conjI)
@@ -724,7 +715,7 @@
apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
apply(simp add:snd_lift)
apply(erule mp)
- apply(force elim:etran.elims intro:Env simp add:lift_def)
+ apply(force elim:etranE intro:Env simp add:lift_def)
apply(simp add:comm_def)
apply(rule conjI)
apply clarify
@@ -766,7 +757,7 @@
back
apply(simp add:snd_lift)
apply(erule mp)
- apply(force elim:etran.elims intro:Env simp add:lift_def)
+ apply(force elim:etranE intro:Env simp add:lift_def)
apply simp
apply clarify
apply(erule_tac x="snd(xs!m)" in allE)
@@ -786,7 +777,7 @@
apply (case_tac i, (simp add:snd_lift)+)
apply(erule mp)
apply(case_tac "xs!m")
- apply(force elim:etran.elims intro:Env simp add:lift_def)
+ apply(force elim:etran.cases intro:Env simp add:lift_def)
apply simp
apply simp
apply clarify
@@ -866,7 +857,7 @@
apply simp
apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps)
apply(erule mp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(case_tac "fst(((Some P, s) # xs) ! i)")
apply(force intro:Env simp add:lift_def)
apply(force intro:Env simp add:lift_def)
@@ -900,7 +891,7 @@
apply(erule mp)
apply(erule tl_of_assum_in_assum,simp)
--{* While-None *}
-apply(ind_cases "((Some (While b P), s), None, t) \<in> ctran")
+apply(ind_cases "((Some (While b P), s), None, t) \<in> ctran" for s t)
apply(simp add:comm_def)
apply(simp add:cptn_iff_cptn_mod [THEN sym])
apply(rule conjI,clarify)
@@ -909,7 +900,7 @@
apply(rule conjI, clarify)
apply(case_tac i,simp,simp)
apply(force simp add:not_ctran_None2)
-apply(subgoal_tac "\<forall>i. Suc i < length ((None, sa) # xs) \<longrightarrow> (((None, sa) # xs) ! i, ((None, sa) # xs) ! Suc i)\<in> etran")
+apply(subgoal_tac "\<forall>i. Suc i < length ((None, t) # xs) \<longrightarrow> (((None, t) # xs) ! i, ((None, t) # xs) ! Suc i)\<in> etran")
prefer 2
apply clarify
apply(rule_tac m="length ((None, s) # xs)" in etran_or_ctran,simp+)
@@ -934,7 +925,7 @@
apply(case_tac "fst(((Some P, sa) # xs) ! i)")
apply(case_tac "((Some P, sa) # xs) ! i")
apply (simp add:lift_def)
- apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t")
+ apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t" for ba t)
apply simp
apply simp
apply(simp add:snd_lift del:map.simps)
@@ -946,9 +937,9 @@
apply(erule_tac x="Suc ia" in allE,simp add:snd_lift del:map.simps)
apply(erule mp)
apply(case_tac "fst(((Some P, sa) # xs) ! ia)")
- apply(erule etran.elims,simp add:lift_def)
+ apply(erule etranE,simp add:lift_def)
apply(rule Env)
- apply(erule etran.elims,simp add:lift_def)
+ apply(erule etranE,simp add:lift_def)
apply(rule Env)
apply (simp add:comm_def del:map.simps)
apply clarify
@@ -986,7 +977,7 @@
apply(case_tac "fst(((Some P, sa) # xs) ! i)")
apply(case_tac "((Some P, sa) # xs) ! i")
apply (simp add:lift_def del:last.simps)
- apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t")
+ apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t" for ba t)
apply simp
apply simp
apply(simp add:snd_lift del:map.simps last.simps)
@@ -998,9 +989,9 @@
apply clarify
apply(erule_tac x="Suc ia" in allE,simp add:nth_append snd_lift del:map.simps last.simps, erule mp)
apply(case_tac "fst(((Some P, sa) # xs) ! ia)")
- apply(erule etran.elims,simp add:lift_def)
+ apply(erule etranE,simp add:lift_def)
apply(rule Env)
- apply(erule etran.elims,simp add:lift_def)
+ apply(erule etranE,simp add:lift_def)
apply(rule Env)
apply (simp add:comm_def del:map.simps)
apply clarify
@@ -1158,9 +1149,9 @@
--{* a c-tran in some @{text "\<sigma>_{ib}"} *}
apply clarify
apply(case_tac "i=ib",simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule_tac x="ib" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE)
- apply (erule etran.elims)
+ apply (erule etranE)
apply(case_tac "ia=m",simp)
apply simp
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (\<forall> i. ?P i j)" in allE)
@@ -1198,7 +1189,7 @@
apply(force simp add:same_state_def par_assum_def)
apply clarify
apply(case_tac "i=ia",simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule_tac x="ia" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE,simp)
apply(erule_tac x=j and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
@@ -1237,7 +1228,7 @@
apply clarify
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in all_dupE)
apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
-apply(erule par_ctran.elims,simp)
+apply(erule par_ctranE,simp)
apply(erule_tac x=i and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
apply(rule_tac x=ia in exI)
@@ -1255,7 +1246,7 @@
done
lemma parcptn_not_empty [simp]:"[] \<notin> par_cptn"
-apply(force elim:par_cptn.elims)
+apply(force elim:par_cptn.cases)
done
lemma five:
@@ -1336,12 +1327,12 @@
apply(case_tac list,simp,simp)
apply(case_tac i)
apply(simp add:par_cp_def ParallelCom_def)
- apply(erule par_ctran.elims,simp)
+ apply(erule par_ctranE,simp)
apply(simp add:par_cp_def ParallelCom_def)
apply clarify
-apply(erule par_cptn.elims,simp)
+apply(erule par_cptn.cases,simp)
apply simp
-apply(erule par_ctran.elims)
+apply(erule par_ctranE)
back
apply simp
done
--- a/src/HOL/HoareParallel/RG_Tran.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/HoareParallel/RG_Tran.thy Wed Jul 11 11:14:51 2007 +0200
@@ -10,73 +10,76 @@
types 'a conf = "(('a com) option) \<times> 'a"
-consts etran :: "('a conf \<times> 'a conf) set"
-syntax "_etran" :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -e\<rightarrow> _" [81,81] 80)
-translations "P -e\<rightarrow> Q" \<rightleftharpoons> "(P,Q) \<in> etran"
-inductive etran
-intros
- Env: "(P, s) -e\<rightarrow> (P, t)"
+inductive_set
+ etran :: "('a conf \<times> 'a conf) set"
+ and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -e\<rightarrow> _" [81,81] 80)
+where
+ "P -e\<rightarrow> Q \<equiv> (P,Q) \<in> etran"
+| Env: "(P, s) -e\<rightarrow> (P, t)"
+
+lemma etranE: "c -e\<rightarrow> c' \<Longrightarrow> (\<And>P s t. c = (P, s) \<Longrightarrow> c' = (P, t) \<Longrightarrow> Q) \<Longrightarrow> Q"
+ by (induct c, induct c', erule etran.cases, blast)
subsubsection {* Component transitions *}
-consts ctran :: "('a conf \<times> 'a conf) set"
-syntax
- "_ctran" :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c\<rightarrow> _" [81,81] 80)
- "_ctran_*":: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c*\<rightarrow> _" [81,81] 80)
-translations
- "P -c\<rightarrow> Q" \<rightleftharpoons> "(P,Q) \<in> ctran"
- "P -c*\<rightarrow> Q" \<rightleftharpoons> "(P,Q) \<in> ctran^*"
+inductive_set
+ ctran :: "('a conf \<times> 'a conf) set"
+ and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c\<rightarrow> _" [81,81] 80)
+ and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c*\<rightarrow> _" [81,81] 80)
+where
+ "P -c\<rightarrow> Q \<equiv> (P,Q) \<in> ctran"
+| "P -c*\<rightarrow> Q \<equiv> (P,Q) \<in> ctran^*"
-inductive ctran
-intros
- Basic: "(Some(Basic f), s) -c\<rightarrow> (None, f s)"
+| Basic: "(Some(Basic f), s) -c\<rightarrow> (None, f s)"
- Seq1: "(Some P0, s) -c\<rightarrow> (None, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some P1, t)"
+| Seq1: "(Some P0, s) -c\<rightarrow> (None, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some P1, t)"
- Seq2: "(Some P0, s) -c\<rightarrow> (Some P2, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some(Seq P2 P1), t)"
+| Seq2: "(Some P0, s) -c\<rightarrow> (Some P2, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some(Seq P2 P1), t)"
- CondT: "s\<in>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P1, s)"
- CondF: "s\<notin>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P2, s)"
+| CondT: "s\<in>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P1, s)"
+| CondF: "s\<notin>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P2, s)"
- WhileF: "s\<notin>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (None, s)"
- WhileT: "s\<in>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (Some(Seq P (While b P)), s)"
+| WhileF: "s\<notin>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (None, s)"
+| WhileT: "s\<in>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (Some(Seq P (While b P)), s)"
- Await: "\<lbrakk>s\<in>b; (Some P, s) -c*\<rightarrow> (None, t)\<rbrakk> \<Longrightarrow> (Some(Await b P), s) -c\<rightarrow> (None, t)"
+| Await: "\<lbrakk>s\<in>b; (Some P, s) -c*\<rightarrow> (None, t)\<rbrakk> \<Longrightarrow> (Some(Await b P), s) -c\<rightarrow> (None, t)"
monos "rtrancl_mono"
subsection {* Semantics of Parallel Programs *}
types 'a par_conf = "('a par_com) \<times> 'a"
-consts
+
+inductive_set
par_etran :: "('a par_conf \<times> 'a par_conf) set"
+ and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
+where
+ "P -pe\<rightarrow> Q \<equiv> (P,Q) \<in> par_etran"
+| ParEnv: "(Ps, s) -pe\<rightarrow> (Ps, t)"
+
+inductive_set
par_ctran :: "('a par_conf \<times> 'a par_conf) set"
-syntax
- "_par_etran":: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
- "_par_ctran":: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
-translations
- "P -pe\<rightarrow> Q" \<rightleftharpoons> "(P,Q) \<in> par_etran"
- "P -pc\<rightarrow> Q" \<rightleftharpoons> "(P,Q) \<in> par_ctran"
+ and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
+where
+ "P -pc\<rightarrow> Q \<equiv> (P,Q) \<in> par_ctran"
+| ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"
-inductive par_etran
-intros
- ParEnv: "(Ps, s) -pe\<rightarrow> (Ps, t)"
-
-inductive par_ctran
-intros
- ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"
+lemma par_ctranE: "c -pc\<rightarrow> c' \<Longrightarrow>
+ (\<And>i Ps s r t. c = (Ps, s) \<Longrightarrow> c' = (Ps[i := r], t) \<Longrightarrow> i < length Ps \<Longrightarrow>
+ (Ps ! i, s) -c\<rightarrow> (r, t) \<Longrightarrow> P) \<Longrightarrow> P"
+ by (induct c, induct c', erule par_ctran.cases, blast)
subsection {* Computations *}
subsubsection {* Sequential computations *}
types 'a confs = "('a conf) list"
-consts cptn :: "('a confs) set"
-inductive "cptn"
-intros
+
+inductive_set cptn :: "('a confs) set"
+where
CptnOne: "[(P,s)] \<in> cptn"
- CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
- CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
+| CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
+| CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
constdefs
cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set"
@@ -85,12 +88,12 @@
subsubsection {* Parallel computations *}
types 'a par_confs = "('a par_conf) list"
-consts par_cptn :: "('a par_confs) set"
-inductive "par_cptn"
-intros
+
+inductive_set par_cptn :: "('a par_confs) set"
+where
ParCptnOne: "[(P,s)] \<in> par_cptn"
- ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
- ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
+| ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
+| ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
constdefs
par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set"
@@ -102,25 +105,24 @@
lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf"
"lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"
-consts cptn_mod :: "('a confs) set"
-inductive "cptn_mod"
-intros
+inductive_set cptn_mod :: "('a confs) set"
+where
CptnModOne: "[(P, s)] \<in> cptn_mod"
- CptnModEnv: "(P, t)#xs \<in> cptn_mod \<Longrightarrow> (P, s)#(P, t)#xs \<in> cptn_mod"
- CptnModNone: "\<lbrakk>(Some P, s) -c\<rightarrow> (None, t); (None, t)#xs \<in> cptn_mod \<rbrakk> \<Longrightarrow> (Some P,s)#(None, t)#xs \<in>cptn_mod"
- CptnModCondT: "\<lbrakk>(Some P0, s)#ys \<in> cptn_mod; s \<in> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P0, s)#ys \<in> cptn_mod"
- CptnModCondF: "\<lbrakk>(Some P1, s)#ys \<in> cptn_mod; s \<notin> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P1, s)#ys \<in> cptn_mod"
- CptnModSeq1: "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; zs=map (lift P1) xs \<rbrakk>
+| CptnModEnv: "(P, t)#xs \<in> cptn_mod \<Longrightarrow> (P, s)#(P, t)#xs \<in> cptn_mod"
+| CptnModNone: "\<lbrakk>(Some P, s) -c\<rightarrow> (None, t); (None, t)#xs \<in> cptn_mod \<rbrakk> \<Longrightarrow> (Some P,s)#(None, t)#xs \<in>cptn_mod"
+| CptnModCondT: "\<lbrakk>(Some P0, s)#ys \<in> cptn_mod; s \<in> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P0, s)#ys \<in> cptn_mod"
+| CptnModCondF: "\<lbrakk>(Some P1, s)#ys \<in> cptn_mod; s \<notin> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P1, s)#ys \<in> cptn_mod"
+| CptnModSeq1: "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; zs=map (lift P1) xs \<rbrakk>
\<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
- CptnModSeq2:
+| CptnModSeq2:
"\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; fst(last ((Some P0, s)#xs)) = None;
(Some P1, snd(last ((Some P0, s)#xs)))#ys \<in> cptn_mod;
zs=(map (lift P1) xs)@ys \<rbrakk> \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
- CptnModWhile1:
+| CptnModWhile1:
"\<lbrakk> (Some P, s)#xs \<in> cptn_mod; s \<in> b; zs=map (lift (While b P)) xs \<rbrakk>
\<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> cptn_mod"
- CptnModWhile2:
+| CptnModWhile2:
"\<lbrakk> (Some P, s)#xs \<in> cptn_mod; fst(last ((Some P, s)#xs))=None; s \<in> b;
zs=(map (lift (While b P)) xs)@ys;
(Some(While b P), snd(last ((Some P, s)#xs)))#ys \<in> cptn_mod\<rbrakk>
@@ -169,7 +171,7 @@
apply simp
apply(simp add:lift_def)
apply clarify
- apply(erule ctran.elims,simp_all)
+ apply(erule ctran.cases,simp_all)
apply clarify
apply(rule_tac x="xs" in exI)
apply simp
@@ -185,10 +187,10 @@
apply simp_all
--{* basic *}
apply clarify
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,rule Basic,simp)
apply clarify
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
--{* Seq1 *}
apply(rule_tac xs="[(None,ta)]" in CptnModSeq2)
apply(erule CptnModNone)
@@ -216,12 +218,12 @@
apply(simp add:lift_def)
--{* Cond *}
apply clarify
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
apply(force elim: CptnModCondT)
apply(force elim: CptnModCondF)
--{* While *}
apply clarify
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule WhileF,simp)
apply(drule div_seq,force)
apply clarify
@@ -231,7 +233,7 @@
apply(force simp add:last_length elim:CptnModWhile2)
--{* await *}
apply clarify
-apply(erule ctran.elims,simp_all)
+apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule Await,simp+)
done
@@ -241,7 +243,7 @@
apply(erule CptnModEnv)
apply(case_tac P)
apply simp
- apply(erule ctran.elims,simp_all)
+ apply(erule ctran.cases,simp_all)
apply(force elim:cptn_onlyif_cptn_mod_aux)
done
@@ -249,7 +251,7 @@
apply(erule cptn.induct)
apply(force simp add:lift_def CptnOne)
apply(force intro:CptnEnv simp add:lift_def)
-apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.elims)
+apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.cases)
done
lemma cptn_append_is_cptn [rule_format]:
@@ -257,7 +259,7 @@
apply(induct c1)
apply simp
apply clarify
-apply(erule cptn.elims,simp_all)
+apply(erule cptn.cases,simp_all)
apply(force intro:CptnEnv)
apply(force elim:CptnComp)
done
@@ -309,7 +311,7 @@
apply(rule CptnComp)
apply(erule CondF,simp)
--{* Seq1 *}
-apply(erule cptn.elims,simp_all)
+apply(erule cptn.cases,simp_all)
apply(rule CptnOne)
apply clarify
apply(drule_tac P=P1 in lift_is_cptn)
@@ -495,10 +497,10 @@
seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym]
if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym]
-lemma prog_not_eq_in_ctran_aux [rule_format]: "(P,s) -c\<rightarrow> (Q,t) \<Longrightarrow> (P\<noteq>Q)"
-apply(erule ctran.induct)
-apply simp_all
-done
+lemma prog_not_eq_in_ctran_aux:
+ assumes c: "(P,s) -c\<rightarrow> (Q,t)"
+ shows "P\<noteq>Q" using c
+ by (induct x1 \<equiv> "(P,s)" x2 \<equiv> "(Q,t)" arbitrary: P s Q t) auto
lemma prog_not_eq_in_ctran [simp]: "\<not> (P,s) -c\<rightarrow> (P,t)"
apply clarify
@@ -522,7 +524,7 @@
done
lemma tl_in_cptn: "\<lbrakk> a#xs \<in>cptn; xs\<noteq>[] \<rbrakk> \<Longrightarrow> xs\<in>cptn"
-apply(force elim:cptn.elims)
+apply(force elim:cptn.cases)
done
lemma tl_zero[rule_format]:
@@ -562,7 +564,7 @@
apply(case_tac "i=ia",simp,simp)
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
apply(drule_tac t=i in not_sym,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(rule ParCptnComp)
apply(erule ParComp,simp)
--{* applying the induction hypothesis *}
@@ -584,7 +586,7 @@
erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
apply(drule_tac t=i in not_sym,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule allE,erule impE,assumption,erule tl_in_cptn)
apply(force simp add:same_length_def length_Suc_conv)
apply(simp add:same_length_def same_state_def)
@@ -620,7 +622,7 @@
apply(rule tl_zero)
apply(erule_tac x=l in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
- apply(force elim:etran.elims intro:Env)
+ apply(force elim:etranE intro:Env)
apply force
apply force
apply simp
@@ -637,7 +639,7 @@
erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
apply(drule_tac t=i in not_sym,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule tl_zero)
apply force
apply force
@@ -654,7 +656,7 @@
apply(erule_tac x=l in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(rule tl_zero)
apply force
apply force
@@ -668,7 +670,7 @@
apply(erule_tac x=ia in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
- apply(force elim:etran.elims intro:Env)
+ apply(force elim:etranE intro:Env)
apply force
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
apply simp
@@ -681,7 +683,7 @@
apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
--{* first step is an environmental step *}
apply clarify
-apply(erule par_etran.elims)
+apply(erule par_etran.cases)
apply simp
apply(rule ParCptnEnv)
apply(erule_tac x="Ps" in allE)
@@ -691,14 +693,14 @@
apply(rule conjI)
apply clarify
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (?I ?s j) \<in> cptn" in allE,simp)
- apply(erule cptn.elims)
+ apply(erule cptn.cases)
apply(simp add:same_length_def)
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
apply(simp add:same_state_def)
apply(erule_tac x=i in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<in>etran" in allE,simp)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(simp add:same_state_def same_length_def)
apply(rule conjI,clarify)
apply(case_tac j,simp,simp)
@@ -725,7 +727,7 @@
apply(rule_tac x=i in exI,simp)
apply(rule conjI)
apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule_tac x=i in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
apply(rule nth_tl_if)
@@ -735,7 +737,7 @@
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
apply clarify
apply(erule_tac x=l and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
- apply(erule etran.elims,simp)
+ apply(erule etranE,simp)
apply(erule_tac x=l in allE, erule impE, assumption,
erule_tac x=1 and P="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
apply(rule nth_tl_if)
@@ -751,7 +753,7 @@
apply(rule tl_zero)
apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
- apply(force elim:etran.elims intro:Env)
+ apply(force elim:etranE intro:Env)
apply force
apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
apply simp
@@ -778,7 +780,7 @@
apply(rule nth_equalityI,simp,simp)
apply(force intro: cptn.intros)
apply(clarify)
-apply(erule par_cptn.elims,simp)
+apply(erule par_cptn.cases,simp)
apply simp
apply(erule_tac x="xs" in allE)
apply(erule_tac x="t" in allE,simp)
@@ -811,7 +813,7 @@
apply clarify
apply(rule_tac x=i in exI,simp)
apply force
-apply(erule par_ctran.elims,simp)
+apply(erule par_ctran.cases,simp)
apply(erule_tac x="Ps[i:=r]" in allE)
apply(erule_tac x="ta" in allE,simp)
apply clarify
@@ -887,7 +889,7 @@
apply(clarify)
apply(simp add:par_cp_def cp_def)
apply(case_tac x)
- apply(force elim:par_cptn.elims)
+ apply(force elim:par_cptn.cases)
apply simp
apply(erule_tac x="list" in allE)
apply clarify
@@ -899,7 +901,7 @@
apply(erule_tac x=0 in allE)
apply(simp add:cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def)
apply clarify
- apply(erule cptn.elims,force,force,force)
+ apply(erule cptn.cases,force,force,force)
apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def)
apply clarify
apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in all_dupE)
--- a/src/HOL/IMP/Compiler.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Compiler.thy Wed Jul 11 11:14:51 2007 +0200
@@ -54,7 +54,7 @@
text {* The other direction! *}
-inductive_cases [elim!]: "(([],p,s),next) : stepa1"
+inductive_cases [elim!]: "(([],p,s),(is',p',s')) : stepa1"
lemma [simp]: "(\<langle>[],q,s\<rangle> -n\<rightarrow> \<langle>p',q',t\<rangle>) = (n=0 \<and> p' = [] \<and> q' = q \<and> t = s)"
apply(rule iffI)
--- a/src/HOL/IMP/Compiler0.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Compiler0.thy Wed Jul 11 11:14:51 2007 +0200
@@ -20,48 +20,32 @@
text {* We describe execution of programs in the machine by
an operational (small step) semantics:
*}
-consts stepa1 :: "instr list \<Rightarrow> ((state\<times>nat) \<times> (state\<times>nat))set"
-
-syntax
- "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ |- (3<_,_>/ -1-> <_,_>)" [50,0,0,0,0] 50)
- "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ |-/ (3<_,_>/ -*-> <_,_>)" [50,0,0,0,0] 50)
-
- "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
- ("_ |-/ (3<_,_>/ -(_)-> <_,_>)" [50,0,0,0,0,0] 50)
-
-syntax (xsymbols)
- "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ \<turnstile> (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
- "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
- "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
- ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
-syntax (HTML output)
- "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ |- (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
- "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
- ("_ |-/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
- "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
- ("_ |-/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
+inductive_set
+ stepa1 :: "instr list \<Rightarrow> ((state\<times>nat) \<times> (state\<times>nat))set"
+ and stepa1' :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
+ ("_ \<turnstile> (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
+ for P :: "instr list"
+where
+ "P \<turnstile> \<langle>s,m\<rangle> -1\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : stepa1 P"
+| ASIN[simp]:
+ "\<lbrakk> n<size P; P!n = ASIN x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s[x\<mapsto> a s],Suc n\<rangle>"
+| JMPFT[simp,intro]:
+ "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,Suc n\<rangle>"
+| JMPFF[simp,intro]:
+ "\<lbrakk> n<size P; P!n = JMPF b i; ~b s; m=n+i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,m\<rangle>"
+| JMPB[simp]:
+ "\<lbrakk> n<size P; P!n = JMPB i; i <= n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,j\<rangle>"
-translations
- "P \<turnstile> \<langle>s,m\<rangle> -1\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : stepa1 P"
- "P \<turnstile> \<langle>s,m\<rangle> -*\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^*)"
- "P \<turnstile> \<langle>s,m\<rangle> -(i)\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^i)"
+abbreviation
+ stepa :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
+ ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50) where
+ "P \<turnstile> \<langle>s,m\<rangle> -*\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : ((stepa1 P)^*)"
-inductive "stepa1 P"
-intros
-ASIN[simp]:
- "\<lbrakk> n<size P; P!n = ASIN x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s[x\<mapsto> a s],Suc n\<rangle>"
-JMPFT[simp,intro]:
- "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,Suc n\<rangle>"
-JMPFF[simp,intro]:
- "\<lbrakk> n<size P; P!n = JMPF b i; ~b s; m=n+i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,m\<rangle>"
-JMPB[simp]:
- "\<lbrakk> n<size P; P!n = JMPB i; i <= n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,j\<rangle>"
+abbreviation
+ stepan :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
+ ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50) where
+ "P \<turnstile> \<langle>s,m\<rangle> -(i)\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : ((stepa1 P)^i)"
subsection "The compiler"
--- a/src/HOL/IMP/Denotation.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Denotation.thy Wed Jul 11 11:14:51 2007 +0200
@@ -62,7 +62,7 @@
apply fast
(* while *)
-apply (erule lfp_induct_set [OF _ Gamma_mono])
+apply (erule lfp_induct2 [OF _ Gamma_mono])
apply (unfold Gamma_def)
apply fast
done
--- a/src/HOL/IMP/Expr.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Expr.thy Wed Jul 11 11:14:51 2007 +0200
@@ -26,16 +26,14 @@
| Op2 "nat => nat => nat" aexp aexp
subsection "Evaluation of arithmetic expressions"
-consts evala :: "((aexp*state) * nat) set"
-syntax "_evala" :: "[aexp*state,nat] => bool" (infixl "-a->" 50)
-translations
- "aesig -a-> n" == "(aesig,n) : evala"
-inductive evala
- intros
+
+inductive
+ evala :: "[aexp*state,nat] => bool" (infixl "-a->" 50)
+where
N: "(N(n),s) -a-> n"
- X: "(X(x),s) -a-> s(x)"
- Op1: "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)"
- Op2: "[| (e0,s) -a-> n0; (e1,s) -a-> n1 |]
+| X: "(X(x),s) -a-> s(x)"
+| Op1: "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)"
+| Op2: "[| (e0,s) -a-> n0; (e1,s) -a-> n1 |]
==> (Op2 f e0 e1,s) -a-> f n0 n1"
lemmas [intro] = N X Op1 Op2
@@ -52,23 +50,19 @@
| ori bexp bexp (infixl "ori" 60)
subsection "Evaluation of boolean expressions"
-consts evalb :: "((bexp*state) * bool)set"
-syntax "_evalb" :: "[bexp*state,bool] => bool" (infixl "-b->" 50)
-translations
- "besig -b-> b" == "(besig,b) : evalb"
-
-inductive evalb
+inductive
+ evalb :: "[bexp*state,bool] => bool" (infixl "-b->" 50)
-- "avoid clash with ML constructors true, false"
- intros
+where
tru: "(true,s) -b-> True"
- fls: "(false,s) -b-> False"
- ROp: "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |]
+| fls: "(false,s) -b-> False"
+| ROp: "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |]
==> (ROp f a0 a1,s) -b-> f n0 n1"
- noti: "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)"
- andi: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
+| noti: "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)"
+| andi: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
==> (b0 andi b1,s) -b-> (w0 & w1)"
- ori: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
+| ori: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
==> (b0 ori b1,s) -b-> (w0 | w1)"
lemmas [intro] = tru fls ROp noti andi ori
@@ -117,21 +111,21 @@
lemma [simp]:
"((ROp f a0 a1,sigma) -b-> w) =
(? m. (a0,sigma) -a-> m & (? n. (a1,sigma) -a-> n & w = f m n))"
- by (rule,cases set: evalb) auto
+ by (rule,cases set: evalb) blast+
lemma [simp]:
"((noti(b),sigma) -b-> w) = (? x. (b,sigma) -b-> x & w = (~x))"
- by (rule,cases set: evalb) auto
+ by (rule,cases set: evalb) blast+
lemma [simp]:
"((b0 andi b1,sigma) -b-> w) =
(? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x&y)))"
- by (rule,cases set: evalb) auto
+ by (rule,cases set: evalb) blast+
lemma [simp]:
"((b0 ori b1,sigma) -b-> w) =
(? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))"
- by (rule,cases set: evalb) auto
+ by (rule,cases set: evalb) blast+
lemma aexp_iff: "((a,s) -a-> n) = (A a s = n)"
--- a/src/HOL/IMP/Hoare.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Hoare.thy Wed Jul 11 11:14:51 2007 +0200
@@ -13,20 +13,17 @@
constdefs hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50)
"|= {P}c{Q} == !s t. (s,t) : C(c) --> P s --> Q t"
-consts hoare :: "(assn * com * assn) set"
-syntax "_hoare" :: "[bool,com,bool] => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
-translations "|- {P}c{Q}" == "(P,c,Q) : hoare"
-
-inductive hoare
-intros
+inductive
+ hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50)
+where
skip: "|- {P}\<SKIP>{P}"
- ass: "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
- semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
- If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==>
+| ass: "|- {%s. P(s[x\<mapsto>a s])} x:==a {P}"
+| semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}"
+| If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==>
|- {P} \<IF> b \<THEN> c \<ELSE> d {Q}"
- While: "|- {%s. P s & b s} c {P} ==>
+| While: "|- {%s. P s & b s} c {P} ==>
|- {P} \<WHILE> b \<DO> c {%s. P s & ~b s}"
- conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==>
+| conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==>
|- {P'}c{Q'}"
constdefs wp :: "com => assn => assn"
--- a/src/HOL/IMP/Machines.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Machines.thy Wed Jul 11 11:14:51 2007 +0200
@@ -36,43 +36,28 @@
subsection "M0 with PC"
-consts exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set"
-syntax
- "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1<_,/_>)/ -1-> (1<_,/_>))" [50,0,0,0,0] 50)
- "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1<_,/_>)/ -*-> (1<_,/_>))" [50,0,0,0,0] 50)
- "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1<_,/_>)/ -_-> (1<_,/_>))" [50,0,0,0,0] 50)
-
-syntax (xsymbols)
- "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
- "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
- "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
- ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
+inductive_set
+ exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set"
+ and exec01' :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
+ ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
+ for P :: "instr list"
+where
+ "p \<turnstile> \<langle>i,s\<rangle> -1\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)"
+| SET: "\<lbrakk> n<size P; P!n = SET x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s[x\<mapsto> a s]\<rangle>"
+| JMPFT: "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s\<rangle>"
+| JMPFF: "\<lbrakk> n<size P; P!n = JMPF b i; \<not>b s; m=n+i+1; m \<le> size P \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>m,s\<rangle>"
+| JMPB: "\<lbrakk> n<size P; P!n = JMPB i; i \<le> n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>j,s\<rangle>"
-syntax (HTML output)
- "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
- "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
- "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
- ("(_/ |- (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
+abbreviation
+ exec0s :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
+ ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50) where
+ "p \<turnstile> \<langle>i,s\<rangle> -*\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)^*"
-translations
- "p \<turnstile> \<langle>i,s\<rangle> -1\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)"
- "p \<turnstile> \<langle>i,s\<rangle> -*\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)^*"
- "p \<turnstile> \<langle>i,s\<rangle> -n\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)^n"
-
-inductive "exec01 P"
-intros
-SET: "\<lbrakk> n<size P; P!n = SET x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s[x\<mapsto> a s]\<rangle>"
-JMPFT: "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s\<rangle>"
-JMPFF: "\<lbrakk> n<size P; P!n = JMPF b i; \<not>b s; m=n+i+1; m \<le> size P \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>m,s\<rangle>"
-JMPB: "\<lbrakk> n<size P; P!n = JMPB i; i \<le> n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>j,s\<rangle>"
+abbreviation
+ exec0n :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
+ ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50) where
+ "p \<turnstile> \<langle>i,s\<rangle> -n\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)^n"
subsection "M0 with lists"
@@ -82,40 +67,31 @@
types config = "instrs \<times> instrs \<times> state"
-consts stepa1 :: "(config \<times> config)set"
-syntax
- "_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
- ("((1<_,/_,/_>)/ -1-> (1<_,/_,/_>))" 50)
- "_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
- ("((1<_,/_,/_>)/ -*-> (1<_,/_,/_>))" 50)
- "_stepan" :: "[state,instrs,instrs, nat, instrs,instrs,state] \<Rightarrow> bool"
- ("((1<_,/_,/_>)/ -_-> (1<_,/_,/_>))" 50)
-
-syntax (xsymbols)
- "_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
- ("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
- "_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
- ("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
- "_stepan" :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool"
- ("((1\<langle>_,/_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
-
-translations
- "\<langle>p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : stepa1"
- "\<langle>p,q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : (stepa1^*)"
- "\<langle>p,q,s\<rangle> -i\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : (stepa1^i)"
-
-
-inductive "stepa1"
-intros
- "\<langle>SET x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,SET x a#q,s[x\<mapsto> a s]\<rangle>"
- "b s \<Longrightarrow> \<langle>JMPF b i#p,q,s\<rangle> -1\<rightarrow> \<langle>p,JMPF b i#q,s\<rangle>"
- "\<lbrakk> \<not> b s; i \<le> size p \<rbrakk>
+inductive_set
+ stepa1 :: "(config \<times> config)set"
+ and stepa1' :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
+ ("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
+where
+ "\<langle>p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : stepa1"
+| "\<langle>SET x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,SET x a#q,s[x\<mapsto> a s]\<rangle>"
+| "b s \<Longrightarrow> \<langle>JMPF b i#p,q,s\<rangle> -1\<rightarrow> \<langle>p,JMPF b i#q,s\<rangle>"
+| "\<lbrakk> \<not> b s; i \<le> size p \<rbrakk>
\<Longrightarrow> \<langle>JMPF b i # p, q, s\<rangle> -1\<rightarrow> \<langle>drop i p, rev(take i p) @ JMPF b i # q, s\<rangle>"
- "i \<le> size q
+| "i \<le> size q
\<Longrightarrow> \<langle>JMPB i # p, q, s\<rangle> -1\<rightarrow> \<langle>rev(take i q) @ JMPB i # p, drop i q, s\<rangle>"
-inductive_cases execE: "((i#is,p,s),next) : stepa1"
+abbreviation
+ stepa :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
+ ("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) where
+ "\<langle>p,q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^*)"
+
+abbreviation
+ stepan :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool"
+ ("((1\<langle>_,/_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) where
+ "\<langle>p,q,s\<rangle> -i\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^i)"
+
+inductive_cases execE: "((i#is,p,s), (is',p',s')) : stepa1"
lemma exec_simp[simp]:
"(\<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>) = (case i of
--- a/src/HOL/IMP/Natural.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Natural.thy Wed Jul 11 11:14:51 2007 +0200
@@ -11,22 +11,12 @@
subsection "Execution of commands"
-consts evalc :: "(com \<times> state \<times> state) set"
-syntax "_evalc" :: "[com,state,state] \<Rightarrow> bool" ("<_,_>/ -c-> _" [0,0,60] 60)
-
-syntax (xsymbols)
- "_evalc" :: "[com,state,state] \<Rightarrow> bool" ("\<langle>_,_\<rangle>/ \<longrightarrow>\<^sub>c _" [0,0,60] 60)
-
-syntax (HTML output)
- "_evalc" :: "[com,state,state] \<Rightarrow> bool" ("\<langle>_,_\<rangle>/ \<longrightarrow>\<^sub>c _" [0,0,60] 60)
-
text {*
We write @{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} for \emph{Statement @{text c}, started
in state @{text s}, terminates in state @{text s'}}. Formally,
@{text "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'"} is just another form of saying \emph{the tuple
@{text "(c,s,s')"} is part of the relation @{text evalc}}:
*}
-translations "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'" == "(c,s,s') \<in> evalc"
constdefs
update :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("_/[_ ::= /_]" [900,0,0] 900)
@@ -38,18 +28,19 @@
text {*
The big-step execution relation @{text evalc} is defined inductively:
*}
-inductive evalc
- intros
+inductive
+ evalc :: "[com,state,state] \<Rightarrow> bool" ("\<langle>_,_\<rangle>/ \<longrightarrow>\<^sub>c _" [0,0,60] 60)
+where
Skip: "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s"
- Assign: "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s[x\<mapsto>a s]"
+| Assign: "\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s[x\<mapsto>a s]"
- Semi: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
+| Semi: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
- IfTrue: "b s \<Longrightarrow> \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
- IfFalse: "\<not>b s \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
+| IfTrue: "b s \<Longrightarrow> \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
+| IfFalse: "\<not>b s \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s' \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s'"
- WhileFalse: "\<not>b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s"
- WhileTrue: "b s \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s'
+| WhileFalse: "\<not>b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s"
+| WhileTrue: "b s \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s'
\<Longrightarrow> \<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s'"
lemmas evalc.intros [intro] -- "use those rules in automatic proofs"
@@ -74,33 +65,33 @@
*}
lemma skip:
"\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma assign:
"\<langle>x :== a,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s[x\<mapsto>a s])"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma semi:
"\<langle>c0; c1, s\<rangle> \<longrightarrow>\<^sub>c s' = (\<exists>s''. \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'' \<and> \<langle>c1,s''\<rangle> \<longrightarrow>\<^sub>c s')"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma ifTrue:
"b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c0,s\<rangle> \<longrightarrow>\<^sub>c s'"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma ifFalse:
"\<not>b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c0 \<ELSE> c1, s\<rangle> \<longrightarrow>\<^sub>c s' = \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c s'"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma whileFalse:
"\<not> b s \<Longrightarrow> \<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>c s' = (s' = s)"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
lemma whileTrue:
"b s \<Longrightarrow>
\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>c s' =
(\<exists>s''. \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c s'' \<and> \<langle>\<WHILE> b \<DO> c, s''\<rangle> \<longrightarrow>\<^sub>c s')"
- by (rule, erule evalc.elims) auto
+ by (rule, erule evalc.cases) auto
text "Again, Isabelle may use these rules in automatic proofs:"
lemmas evalc_cases [simp] = skip assign ifTrue ifFalse whileFalse semi whileTrue
--- a/src/HOL/IMP/Transition.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/Transition.thy Wed Jul 11 11:14:51 2007 +0200
@@ -20,10 +20,8 @@
a statement, the transition relation is not
@{typ "((com \<times> state) \<times> (com \<times> state)) set"}
but instead:
-*}
-consts evalc1 :: "((com option \<times> state) \<times> (com option \<times> state)) set"
+ @{typ "((com option \<times> state) \<times> (com option \<times> state)) set"}
-text {*
Some syntactic sugar that we will use to hide the
@{text option} part in configurations:
*}
@@ -44,53 +42,40 @@
"\<langle>s\<rangle>" == "(None, s)"
text {*
+ Now, finally, we are set to write down the rules for our small step semantics:
+*}
+inductive_set
+ evalc1 :: "((com option \<times> state) \<times> (com option \<times> state)) set"
+ and evalc1' :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
+ ("_ \<longrightarrow>\<^sub>1 _" [60,60] 61)
+where
+ "cs \<longrightarrow>\<^sub>1 cs' == (cs,cs') \<in> evalc1"
+| Skip: "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s\<rangle>"
+| Assign: "\<langle>x :== a, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s[x \<mapsto> a s]\<rangle>"
+
+| Semi1: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s'\<rangle>"
+| Semi2: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0',s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0';c1,s'\<rangle>"
+
+| IfTrue: "b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s\<rangle>"
+| IfFalse: "\<not>b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c2,s\<rangle>"
+
+| While: "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>,s\<rangle>"
+
+lemmas [intro] = evalc1.intros -- "again, use these rules in automatic proofs"
+
+text {*
More syntactic sugar for the transition relation, and its
iteration.
*}
-syntax
- "_evalc1" :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
- ("_ -1-> _" [60,60] 60)
- "_evalcn" :: "[(com option\<times>state),nat,(com option\<times>state)] \<Rightarrow> bool"
- ("_ -_-> _" [60,60,60] 60)
- "_evalc*" :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
- ("_ -*-> _" [60,60] 60)
-
-syntax (xsymbols)
- "_evalc1" :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
- ("_ \<longrightarrow>\<^sub>1 _" [60,60] 61)
- "_evalcn" :: "[(com option\<times>state),nat,(com option\<times>state)] \<Rightarrow> bool"
- ("_ -_\<rightarrow>\<^sub>1 _" [60,60,60] 60)
- "_evalc*" :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
- ("_ \<longrightarrow>\<^sub>1\<^sup>* _" [60,60] 60)
-
-translations
- "cs \<longrightarrow>\<^sub>1 cs'" == "(cs,cs') \<in> evalc1"
- "cs -n\<rightarrow>\<^sub>1 cs'" == "(cs,cs') \<in> evalc1^n"
- "cs \<longrightarrow>\<^sub>1\<^sup>* cs'" == "(cs,cs') \<in> evalc1^*"
+abbreviation
+ evalcn :: "[(com option\<times>state),nat,(com option\<times>state)] \<Rightarrow> bool"
+ ("_ -_\<rightarrow>\<^sub>1 _" [60,60,60] 60) where
+ "cs -n\<rightarrow>\<^sub>1 cs' == (cs,cs') \<in> evalc1^n"
- -- {* Isabelle/HOL converts @{text "(cs0,(c1,s1))"} to @{term "(cs0,c1,s1)"},
- so we also include: *}
- "cs0 \<longrightarrow>\<^sub>1 (c1,s1)" == "(cs0,c1,s1) \<in> evalc1"
- "cs0 -n\<rightarrow>\<^sub>1 (c1,s1)" == "(cs0,c1,s1) \<in> evalc1^n"
- "cs0 \<longrightarrow>\<^sub>1\<^sup>* (c1,s1)" == "(cs0,c1,s1) \<in> evalc1^*"
-
-text {*
- Now, finally, we are set to write down the rules for our small step semantics:
-*}
-inductive evalc1
- intros
- Skip: "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s\<rangle>"
- Assign: "\<langle>x :== a, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s[x \<mapsto> a s]\<rangle>"
-
- Semi1: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s'\<rangle>"
- Semi2: "\<langle>c0,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0',s'\<rangle> \<Longrightarrow> \<langle>c0;c1,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c0';c1,s'\<rangle>"
-
- IfTrue: "b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1,s\<rangle>"
- IfFalse: "\<not>b s \<Longrightarrow> \<langle>\<IF> b \<THEN> c1 \<ELSE> c2,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c2,s\<rangle>"
-
- While: "\<langle>\<WHILE> b \<DO> c,s\<rangle> \<longrightarrow>\<^sub>1 \<langle>\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>,s\<rangle>"
-
-lemmas [intro] = evalc1.intros -- "again, use these rules in automatic proofs"
+abbreviation
+ evalc' :: "[(com option\<times>state),(com option\<times>state)] \<Rightarrow> bool"
+ ("_ \<longrightarrow>\<^sub>1\<^sup>* _" [60,60] 60) where
+ "cs \<longrightarrow>\<^sub>1\<^sup>* cs' == (cs,cs') \<in> evalc1^*"
(*<*)
(* fixme: move to Relation_Power.thy *)
@@ -120,18 +105,18 @@
syntax directed way:
*}
lemma SKIP_1: "\<langle>\<SKIP>, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>s\<rangle>)"
- by (rule, cases set: evalc1, auto)
+ by (induct y, rule, cases set: evalc1, auto)
lemma Assign_1: "\<langle>x :== a, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>s[x \<mapsto> a s]\<rangle>)"
- by (rule, cases set: evalc1, auto)
+ by (induct y, rule, cases set: evalc1, auto)
lemma Cond_1:
"\<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 y = ((b s \<longrightarrow> y = \<langle>c1, s\<rangle>) \<and> (\<not>b s \<longrightarrow> y = \<langle>c2, s\<rangle>))"
- by (rule, cases set: evalc1, auto)
+ by (induct y, rule, cases set: evalc1, auto)
lemma While_1:
"\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 y = (y = \<langle>\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>, s\<rangle>)"
- by (rule, cases set: evalc1, auto)
+ by (induct y, rule, cases set: evalc1, auto)
lemmas [simp] = SKIP_1 Assign_1 Cond_1 While_1
@@ -197,10 +182,10 @@
has terminated and there is no next configuration:
*}
lemma stuck [elim!]: "\<langle>s\<rangle> \<longrightarrow>\<^sub>1 y \<Longrightarrow> P"
- by (auto elim: evalc1.elims)
+ by (induct y, auto elim: evalc1.cases)
lemma evalc_None_retrancl [simp, dest!]: "\<langle>s\<rangle> \<longrightarrow>\<^sub>1\<^sup>* s' \<Longrightarrow> s' = \<langle>s\<rangle>"
- by (induct set: rtrancl_set) auto
+ by (induct set: rtrancl) auto
(*<*)
(* FIXME: relpow.simps don't work *)
@@ -230,10 +215,10 @@
case (Suc n)
from `\<langle>c1; c2, s\<rangle> -Suc n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>`
- obtain y where
- 1: "\<langle>c1; c2, s\<rangle> \<longrightarrow>\<^sub>1 y" and
- n: "y -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>"
- by blast
+ obtain co s''' where
+ 1: "\<langle>c1; c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s''')" and
+ n: "(co, s''') -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>"
+ by auto
from 1
show "\<exists>i j s'. \<langle>c1, s\<rangle> -i\<rightarrow>\<^sub>1 \<langle>s'\<rangle> \<and> \<langle>c2, s'\<rangle> -j\<rightarrow>\<^sub>1 \<langle>s''\<rangle> \<and> Suc n = i+j"
@@ -241,14 +226,14 @@
proof (cases set: evalc1)
case Semi1
then obtain s' where
- "y = \<langle>c2, s'\<rangle>" and "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle>"
+ "co = Some c2" and "s''' = s'" and "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>s'\<rangle>"
by auto
with 1 n have "?Q 1 n s'" by simp
thus ?thesis by blast
next
case Semi2
then obtain c1' s' where
- y: "y = \<langle>c1'; c2, s'\<rangle>" and
+ "co = Some (c1'; c2)" "s''' = s'" and
c1: "\<langle>c1, s\<rangle> \<longrightarrow>\<^sub>1 \<langle>c1', s'\<rangle>"
by auto
with n have "\<langle>c1'; c2, s'\<rangle> -n\<rightarrow>\<^sub>1 \<langle>s''\<rangle>" by simp
@@ -476,13 +461,13 @@
qed
inductive_cases evalc1_SEs:
- "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>1 t"
- "\<langle>x:==a,s\<rangle> \<longrightarrow>\<^sub>1 t"
- "\<langle>c1;c2, s\<rangle> \<longrightarrow>\<^sub>1 t"
- "\<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 t"
- "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 t"
+ "\<langle>\<SKIP>,s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
+ "\<langle>x:==a,s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
+ "\<langle>c1;c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
+ "\<langle>\<IF> b \<THEN> c1 \<ELSE> c2, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
+ "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
-inductive_cases evalc1_E: "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 t"
+inductive_cases evalc1_E: "\<langle>\<WHILE> b \<DO> c, s\<rangle> \<longrightarrow>\<^sub>1 (co, s')"
declare evalc1_SEs [elim!]
@@ -546,6 +531,7 @@
apply (intro strip)
apply (erule rel_pow_E2)
apply simp
+apply (simp only: split_paired_all)
apply (erule evalc1_E)
apply simp
--- a/src/HOL/IMP/VC.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMP/VC.thy Wed Jul 11 11:14:51 2007 +0200
@@ -120,11 +120,11 @@
show ?case (is "? ac. ?C ac")
proof show "?C Askip" by simp qed
next
- case (ass P a x)
+ case (ass P x a)
show ?case (is "? ac. ?C ac")
proof show "?C(Aass x a)" by simp qed
next
- case (semi P Q R c1 c2)
+ case (semi P c1 Q c2 R)
from semi.hyps obtain ac1 where ih1: "?Eq P c1 Q ac1" by fast
from semi.hyps obtain ac2 where ih2: "?Eq Q c2 R ac2" by fast
show ?case (is "? ac. ?C ac")
@@ -133,7 +133,7 @@
using ih1 ih2 by simp (fast elim!: awp_mono vc_mono)
qed
next
- case (If P Q b c1 c2)
+ case (If P b c1 Q c2)
from If.hyps obtain ac1 where ih1: "?Eq (%s. P s & b s) c1 Q ac1" by fast
from If.hyps obtain ac2 where ih2: "?Eq (%s. P s & ~b s) c2 Q ac2" by fast
show ?case (is "? ac. ?C ac")
--- a/src/HOL/IMPP/Com.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMPP/Com.thy Wed Jul 11 11:14:51 2007 +0200
@@ -48,32 +48,29 @@
(* Well-typedness: all procedures called must exist *)
-consts WTs :: "com set"
-syntax WT :: "com => bool"
-translations "WT c" == "c : WTs"
-inductive WTs intros
+inductive WT :: "com => bool" where
Skip: "WT SKIP"
- Assign: "WT (X :== a)"
+ | Assign: "WT (X :== a)"
- Local: "WT c ==>
+ | Local: "WT c ==>
WT (LOCAL Y := a IN c)"
- Semi: "[| WT c0; WT c1 |] ==>
+ | Semi: "[| WT c0; WT c1 |] ==>
WT (c0;; c1)"
- If: "[| WT c0; WT c1 |] ==>
+ | If: "[| WT c0; WT c1 |] ==>
WT (IF b THEN c0 ELSE c1)"
- While: "WT c ==>
+ | While: "WT c ==>
WT (WHILE b DO c)"
- Body: "body pn ~= None ==>
+ | Body: "body pn ~= None ==>
WT (BODY pn)"
- Call: "WT (BODY pn) ==>
+ | Call: "WT (BODY pn) ==>
WT (X:=CALL pn(a))"
inductive_cases WTs_elim_cases:
--- a/src/HOL/IMPP/Hoare.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMPP/Hoare.thy Wed Jul 11 11:14:51 2007 +0200
@@ -34,59 +34,58 @@
consts
triple_valid :: "nat => 'a triple => bool" ( "|=_:_" [0 , 58] 57)
hoare_valids :: "'a triple set => 'a triple set => bool" ("_||=_" [58, 58] 57)
- hoare_derivs :: "('a triple set * 'a triple set) set"
syntax
triples_valid:: "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57)
hoare_valid :: "'a triple set => 'a triple => bool" ("_|=_" [58, 58] 57)
-"@hoare_derivs":: "'a triple set => 'a triple set => bool" ("_||-_" [58, 58] 57)
-"@hoare_deriv" :: "'a triple set => 'a triple => bool" ("_|-_" [58, 58] 57)
defs triple_valid_def: "|=n:t == case t of {P}.c.{Q} =>
!Z s. P Z s --> (!s'. <c,s> -n-> s' --> Q Z s')"
translations "||=n:G" == "Ball G (triple_valid n)"
defs hoare_valids_def: "G||=ts == !n. ||=n:G --> ||=n:ts"
translations "G |=t " == " G||={t}"
- "G||-ts" == "(G,ts) : hoare_derivs"
- "G |-t" == " G||-{t}"
(* Most General Triples *)
constdefs MGT :: "com => state triple" ("{=}._.{->}" [60] 58)
"{=}.c.{->} == {%Z s0. Z = s0}. c .{%Z s1. <c,Z> -c-> s1}"
-inductive hoare_derivs intros
+inductive
+ hoare_derivs :: "'a triple set => 'a triple set => bool" ("_||-_" [58, 58] 57)
+ and hoare_deriv :: "'a triple set => 'a triple => bool" ("_|-_" [58, 58] 57)
+where
+ "G |-t == G||-{t}"
- empty: "G||-{}"
- insert: "[| G |-t; G||-ts |]
+| empty: "G||-{}"
+| insert: "[| G |-t; G||-ts |]
==> G||-insert t ts"
- asm: "ts <= G ==>
+| asm: "ts <= G ==>
G||-ts" (* {P}.BODY pn.{Q} instead of (general) t for SkipD_lemma *)
- cut: "[| G'||-ts; G||-G' |] ==> G||-ts" (* for convenience and efficiency *)
+| cut: "[| G'||-ts; G||-G' |] ==> G||-ts" (* for convenience and efficiency *)
- weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts"
+| weaken: "[| G||-ts' ; ts <= ts' |] ==> G||-ts"
- conseq: "!Z s. P Z s --> (? P' Q'. G|-{P'}.c.{Q'} &
+| 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}"
- Skip: "G|-{P}. SKIP .{P}"
+| Skip: "G|-{P}. SKIP .{P}"
- Ass: "G|-{%Z s. P Z (s[X::=a s])}. X:==a .{P}"
+| Ass: "G|-{%Z s. P Z (s[X::=a s])}. X:==a .{P}"
- Local: "G|-{P}. c .{%Z s. Q Z (s[Loc X::=s'<X>])}
+| Local: "G|-{P}. c .{%Z s. Q Z (s[Loc X::=s'<X>])}
==> G|-{%Z s. s'=s & P Z (s[Loc X::=a s])}. LOCAL X:=a IN c .{Q}"
- Comp: "[| G|-{P}.c.{Q};
+| Comp: "[| G|-{P}.c.{Q};
G|-{Q}.d.{R} |]
==> G|-{P}. (c;;d) .{R}"
- If: "[| G|-{P &> b }.c.{Q};
+| If: "[| G|-{P &> b }.c.{Q};
G|-{P &> (Not o b)}.d.{Q} |]
==> G|-{P}. IF b THEN c ELSE d .{Q}"
- Loop: "G|-{P &> b}.c.{P} ==>
+| Loop: "G|-{P &> b}.c.{P} ==>
G|-{P}. WHILE b DO c .{P &> (Not o b)}"
(*
@@ -94,11 +93,11 @@
|-{P}. the (body pn) .{Q} ==>
G|-{P}. BODY pn .{Q}"
*)
- Body: "[| G Un (%p. {P p}. BODY p .{Q p})`Procs
+| Body: "[| G Un (%p. {P p}. BODY p .{Q p})`Procs
||-(%p. {P p}. the (body p) .{Q p})`Procs |]
==> G||-(%p. {P p}. BODY p .{Q p})`Procs"
- Call: "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])}
+| Call: "G|-{P}. BODY pn .{%Z s. Q Z (setlocs s (getlocs s')[X::=s<Res>])}
==> G|-{%Z s. s'=s & P Z (setlocs s newlocs[Loc Arg::=a s])}.
X:=CALL pn(a) .{Q}"
@@ -283,7 +282,7 @@
apply (blast) (* asm *)
apply (blast) (* cut *)
apply (blast) (* weaken *)
-apply (tactic {* ALLGOALS (EVERY'[REPEAT o thin_tac "?x : hoare_derivs", SIMPSET' simp_tac, CLASET' clarify_tac, REPEAT o smp_tac 1]) *})
+apply (tactic {* ALLGOALS (EVERY'[REPEAT o thin_tac "hoare_derivs ?x ?y", SIMPSET' simp_tac, CLASET' clarify_tac, REPEAT o smp_tac 1]) *})
apply (simp_all (no_asm_use) add: triple_valid_def2)
apply (intro strip, tactic "smp_tac 2 1", blast) (* conseq *)
apply (tactic {* ALLGOALS (CLASIMPSET' clarsimp_tac) *}) (* Skip, Ass, Local *)
--- a/src/HOL/IMPP/Natural.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/IMPP/Natural.thy Wed Jul 11 11:14:51 2007 +0200
@@ -11,17 +11,6 @@
begin
(** Execution of commands **)
-consts
- evalc :: "(com * state * state) set"
- evaln :: "(com * state * nat * state) set"
-
-syntax
- "@evalc":: "[com,state, state] => bool" ("<_,_>/ -c-> _" [0,0, 51] 51)
- "@evaln":: "[com,state,nat,state] => bool" ("<_,_>/ -_-> _" [0,0,0,51] 51)
-
-translations
- "<c,s> -c-> s'" == "(c,s, s') : evalc"
- "<c,s> -n-> s'" == "(c,s,n,s') : evaln"
consts
newlocs :: locals
@@ -33,63 +22,65 @@
translations
"s<X>" == "getlocs s X"
-inductive evalc
- intros
+inductive
+ evalc :: "[com,state, state] => bool" ("<_,_>/ -c-> _" [0,0, 51] 51)
+ where
Skip: "<SKIP,s> -c-> s"
- Assign: "<X :== a,s> -c-> s[X::=a s]"
+ | Assign: "<X :== a,s> -c-> s[X::=a s]"
- Local: "<c, s0[Loc Y::= a s0]> -c-> s1 ==>
+ | Local: "<c, s0[Loc Y::= a s0]> -c-> s1 ==>
<LOCAL Y := a IN c, s0> -c-> s1[Loc Y::=s0<Y>]"
- Semi: "[| <c0,s0> -c-> s1; <c1,s1> -c-> s2 |] ==>
+ | Semi: "[| <c0,s0> -c-> s1; <c1,s1> -c-> s2 |] ==>
<c0;; c1, s0> -c-> s2"
- IfTrue: "[| b s; <c0,s> -c-> s1 |] ==>
+ | IfTrue: "[| b s; <c0,s> -c-> s1 |] ==>
<IF b THEN c0 ELSE c1, s> -c-> s1"
- IfFalse: "[| ~b s; <c1,s> -c-> s1 |] ==>
+ | IfFalse: "[| ~b s; <c1,s> -c-> s1 |] ==>
<IF b THEN c0 ELSE c1, s> -c-> s1"
- WhileFalse: "~b s ==> <WHILE b DO c,s> -c-> s"
+ | WhileFalse: "~b s ==> <WHILE b DO c,s> -c-> s"
- WhileTrue: "[| b s0; <c,s0> -c-> s1; <WHILE b DO c, s1> -c-> s2 |] ==>
+ | WhileTrue: "[| b s0; <c,s0> -c-> s1; <WHILE b DO c, s1> -c-> s2 |] ==>
<WHILE b DO c, s0> -c-> s2"
- Body: "<the (body pn), s0> -c-> s1 ==>
+ | Body: "<the (body pn), s0> -c-> s1 ==>
<BODY pn, s0> -c-> s1"
- Call: "<BODY pn, (setlocs s0 newlocs)[Loc Arg::=a s0]> -c-> s1 ==>
+ | Call: "<BODY pn, (setlocs s0 newlocs)[Loc Arg::=a s0]> -c-> s1 ==>
<X:=CALL pn(a), s0> -c-> (setlocs s1 (getlocs s0))
[X::=s1<Res>]"
-inductive evaln
- intros
+inductive
+ evaln :: "[com,state,nat,state] => bool" ("<_,_>/ -_-> _" [0,0,0,51] 51)
+ where
Skip: "<SKIP,s> -n-> s"
- Assign: "<X :== a,s> -n-> s[X::=a s]"
+ | Assign: "<X :== a,s> -n-> s[X::=a s]"
- Local: "<c, s0[Loc Y::= a s0]> -n-> s1 ==>
+ | Local: "<c, s0[Loc Y::= a s0]> -n-> s1 ==>
<LOCAL Y := a IN c, s0> -n-> s1[Loc Y::=s0<Y>]"
- Semi: "[| <c0,s0> -n-> s1; <c1,s1> -n-> s2 |] ==>
+ | Semi: "[| <c0,s0> -n-> s1; <c1,s1> -n-> s2 |] ==>
<c0;; c1, s0> -n-> s2"
- IfTrue: "[| b s; <c0,s> -n-> s1 |] ==>
+ | IfTrue: "[| b s; <c0,s> -n-> s1 |] ==>
<IF b THEN c0 ELSE c1, s> -n-> s1"
- IfFalse: "[| ~b s; <c1,s> -n-> s1 |] ==>
+ | IfFalse: "[| ~b s; <c1,s> -n-> s1 |] ==>
<IF b THEN c0 ELSE c1, s> -n-> s1"
- WhileFalse: "~b s ==> <WHILE b DO c,s> -n-> s"
+ | WhileFalse: "~b s ==> <WHILE b DO c,s> -n-> s"
- WhileTrue: "[| b s0; <c,s0> -n-> s1; <WHILE b DO c, s1> -n-> s2 |] ==>
+ | WhileTrue: "[| b s0; <c,s0> -n-> s1; <WHILE b DO c, s1> -n-> s2 |] ==>
<WHILE b DO c, s0> -n-> s2"
- Body: "<the (body pn), s0> - n-> s1 ==>
+ | Body: "<the (body pn), s0> - n-> s1 ==>
<BODY pn, s0> -Suc n-> s1"
- Call: "<BODY pn, (setlocs s0 newlocs)[Loc Arg::=a s0]> -n-> s1 ==>
+ | Call: "<BODY pn, (setlocs s0 newlocs)[Loc Arg::=a s0]> -n-> s1 ==>
<X:=CALL pn(a), s0> -n-> (setlocs s1 (getlocs s0))
[X::=s1<Res>]"
--- a/src/HOL/Induct/Com.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/Com.thy Wed Jul 11 11:14:51 2007 +0200
@@ -15,8 +15,6 @@
types state = "loc => nat"
n2n2n = "nat => nat => nat"
-arities loc :: type
-
datatype
exp = N nat
| X loc
@@ -33,36 +31,38 @@
subsection {* Commands *}
text{* Execution of commands *}
-consts exec :: "((exp*state) * (nat*state)) set => ((com*state)*state)set"
-
-abbreviation
- exec_rel ("_/ -[_]-> _" [50,0,50] 50)
- "csig -[eval]-> s == (csig,s) \<in> exec eval"
abbreviation (input)
- generic_rel ("_/ -|[_]-> _" [50,0,50] 50)
+ generic_rel ("_/ -|[_]-> _" [50,0,50] 50) where
"esig -|[eval]-> ns == (esig,ns) \<in> eval"
text{*Command execution. Natural numbers represent Booleans: 0=True, 1=False*}
-inductive "exec eval"
- intros
- Skip: "(SKIP,s) -[eval]-> s"
- Assign: "(e,s) -|[eval]-> (v,s') ==> (x := e, s) -[eval]-> s'(x:=v)"
+inductive_set
+ exec :: "((exp*state) * (nat*state)) set => ((com*state)*state)set"
+ and exec_rel :: "com * state => ((exp*state) * (nat*state)) set => state => bool"
+ ("_/ -[_]-> _" [50,0,50] 50)
+ for eval :: "((exp*state) * (nat*state)) set"
+ where
+ "csig -[eval]-> s == (csig,s) \<in> exec eval"
- Semi: "[| (c0,s) -[eval]-> s2; (c1,s2) -[eval]-> s1 |]
+ | Skip: "(SKIP,s) -[eval]-> s"
+
+ | Assign: "(e,s) -|[eval]-> (v,s') ==> (x := e, s) -[eval]-> s'(x:=v)"
+
+ | Semi: "[| (c0,s) -[eval]-> s2; (c1,s2) -[eval]-> s1 |]
==> (c0 ;; c1, s) -[eval]-> s1"
- IfTrue: "[| (e,s) -|[eval]-> (0,s'); (c0,s') -[eval]-> s1 |]
+ | IfTrue: "[| (e,s) -|[eval]-> (0,s'); (c0,s') -[eval]-> s1 |]
==> (IF e THEN c0 ELSE c1, s) -[eval]-> s1"
- IfFalse: "[| (e,s) -|[eval]-> (Suc 0, s'); (c1,s') -[eval]-> s1 |]
+ | IfFalse: "[| (e,s) -|[eval]-> (Suc 0, s'); (c1,s') -[eval]-> s1 |]
==> (IF e THEN c0 ELSE c1, s) -[eval]-> s1"
- WhileFalse: "(e,s) -|[eval]-> (Suc 0, s1)
+ | WhileFalse: "(e,s) -|[eval]-> (Suc 0, s1)
==> (WHILE e DO c, s) -[eval]-> s1"
- WhileTrue: "[| (e,s) -|[eval]-> (0,s1);
+ | WhileTrue: "[| (e,s) -|[eval]-> (0,s1);
(c,s1) -[eval]-> s2; (WHILE e DO c, s2) -[eval]-> s3 |]
==> (WHILE e DO c, s) -[eval]-> s3"
@@ -79,11 +79,20 @@
text{*Justifies using "exec" in the inductive definition of "eval"*}
lemma exec_mono: "A<=B ==> exec(A) <= exec(B)"
-apply (unfold exec.defs )
-apply (rule lfp_mono)
-apply (assumption | rule basic_monos)+
+apply (rule subsetI)
+apply (simp add: split_paired_all)
+apply (erule exec.induct)
+apply blast+
done
+lemma [pred_set_conv]:
+ "((\<lambda>x x' y y'. ((x, x'), (y, y')) \<in> R) <= (\<lambda>x x' y y'. ((x, x'), (y, y')) \<in> S)) = (R <= S)"
+ by (auto simp add: le_fun_def le_bool_def)
+
+lemma [pred_set_conv]:
+ "((\<lambda>x x' y. ((x, x'), y) \<in> R) <= (\<lambda>x x' y. ((x, x'), y) \<in> S)) = (R <= S)"
+ by (auto simp add: le_fun_def le_bool_def)
+
ML {*
Unify.trace_bound := 30;
Unify.search_bound := 60;
@@ -102,23 +111,21 @@
subsection {* Expressions *}
text{* Evaluation of arithmetic expressions *}
-consts
- eval :: "((exp*state) * (nat*state)) set"
-
-abbreviation
- eval_rel :: "[exp*state,nat*state] => bool" (infixl "-|->" 50)
- "esig -|-> ns == (esig, ns) \<in> eval"
-inductive eval
- intros
- N [intro!]: "(N(n),s) -|-> (n,s)"
+inductive_set
+ eval :: "((exp*state) * (nat*state)) set"
+ and eval_rel :: "[exp*state,nat*state] => bool" (infixl "-|->" 50)
+ where
+ "esig -|-> ns == (esig, ns) \<in> eval"
- X [intro!]: "(X(x),s) -|-> (s(x),s)"
+ | N [intro!]: "(N(n),s) -|-> (n,s)"
- Op [intro]: "[| (e0,s) -|-> (n0,s0); (e1,s0) -|-> (n1,s1) |]
+ | X [intro!]: "(X(x),s) -|-> (s(x),s)"
+
+ | Op [intro]: "[| (e0,s) -|-> (n0,s0); (e1,s0) -|-> (n1,s1) |]
==> (Op f e0 e1, s) -|-> (f n0 n1, s1)"
- valOf [intro]: "[| (c,s) -[eval]-> s0; (e,s0) -|-> (n,s1) |]
+ | valOf [intro]: "[| (c,s) -[eval]-> s0; (e,s0) -|-> (n,s1) |]
==> (VALOF c RESULTIS e, s) -|-> (n, s1)"
monos exec_mono
@@ -135,7 +142,7 @@
by (rule fun_upd_same [THEN subst], fast)
-text{* Make the induction rule look nicer -- though eta_contract makes the new
+text{* Make the induction rule look nicer -- though @{text eta_contract} makes the new
version look worse than it is...*}
lemma split_lemma:
@@ -167,11 +174,11 @@
done
-text{*Lemma for Function_eval. The major premise is that (c,s) executes to s1
+text{*Lemma for @{text Function_eval}. The major premise is that @{text "(c,s)"} executes to @{text "s1"}
using eval restricted to its functional part. Note that the execution
- (c,s) -[eval]-> s2 can use unrestricted eval! The reason is that
- the execution (c,s) -[eval Int {...}]-> s1 assures us that execution is
- functional on the argument (c,s).
+ @{text "(c,s) -[eval]-> s2"} can use unrestricted @{text eval}! The reason is that
+ the execution @{text "(c,s) -[eval Int {...}]-> s1"} assures us that execution is
+ functional on the argument @{text "(c,s)"}.
*}
lemma com_Unique:
"(c,s) -[eval Int {((e,s),(n,t)). \<forall>nt'. (e,s) -|-> nt' --> (n,t)=nt'}]-> s1
--- a/src/HOL/Induct/Comb.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/Comb.thy Wed Jul 11 11:14:51 2007 +0200
@@ -34,47 +34,39 @@
(multi-step) reductions, @{text "--->"}.
*}
-consts
- contract :: "(comb*comb) set"
-
-abbreviation
- contract_rel1 :: "[comb,comb] => bool" (infixl "-1->" 50) where
- "x -1-> y == (x,y) \<in> contract"
+inductive_set
+ contract :: "(comb*comb) set"
+ and contract_rel1 :: "[comb,comb] => bool" (infixl "-1->" 50)
+ where
+ "x -1-> y == (x,y) \<in> contract"
+ | K: "K##x##y -1-> x"
+ | S: "S##x##y##z -1-> (x##z)##(y##z)"
+ | Ap1: "x-1->y ==> x##z -1-> y##z"
+ | Ap2: "x-1->y ==> z##x -1-> z##y"
abbreviation
contract_rel :: "[comb,comb] => bool" (infixl "--->" 50) where
"x ---> y == (x,y) \<in> contract^*"
-inductive contract
- intros
- K: "K##x##y -1-> x"
- S: "S##x##y##z -1-> (x##z)##(y##z)"
- Ap1: "x-1->y ==> x##z -1-> y##z"
- Ap2: "x-1->y ==> z##x -1-> z##y"
-
text {*
Inductive definition of parallel contractions, @{text "=1=>"} and
(multi-step) parallel reductions, @{text "===>"}.
*}
-consts
+inductive_set
parcontract :: "(comb*comb) set"
-
-abbreviation
- parcontract_rel1 :: "[comb,comb] => bool" (infixl "=1=>" 50) where
- "x =1=> y == (x,y) \<in> parcontract"
+ and parcontract_rel1 :: "[comb,comb] => bool" (infixl "=1=>" 50)
+ where
+ "x =1=> y == (x,y) \<in> parcontract"
+ | refl: "x =1=> x"
+ | K: "K##x##y =1=> x"
+ | S: "S##x##y##z =1=> (x##z)##(y##z)"
+ | Ap: "[| x=1=>y; z=1=>w |] ==> x##z =1=> y##w"
abbreviation
parcontract_rel :: "[comb,comb] => bool" (infixl "===>" 50) where
"x ===> y == (x,y) \<in> parcontract^*"
-inductive parcontract
- intros
- refl: "x =1=> x"
- K: "K##x##y =1=> x"
- S: "S##x##y##z =1=> (x##z)##(y##z)"
- Ap: "[| x=1=>y; z=1=>w |] ==> x##z =1=> y##w"
-
text {*
Misc definitions.
*}
--- a/src/HOL/Induct/LFilter.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/LFilter.thy Wed Jul 11 11:14:51 2007 +0200
@@ -9,13 +9,12 @@
theory LFilter imports LList begin
-consts
+inductive_set
findRel :: "('a => bool) => ('a llist * 'a llist)set"
-
-inductive "findRel p"
- intros
+ for p :: "'a => bool"
+ where
found: "p x ==> (LCons x l, LCons x l) \<in> findRel p"
- seek: "[| ~p x; (l,l') \<in> findRel p |] ==> (LCons x l, l') \<in> findRel p"
+ | seek: "[| ~p x; (l,l') \<in> findRel p |] ==> (LCons x l, l') \<in> findRel p"
declare findRel.intros [intro]
--- a/src/HOL/Induct/LList.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/LList.thy Wed Jul 11 11:14:51 2007 +0200
@@ -24,21 +24,19 @@
theory LList imports SList begin
-consts
-
+coinductive_set
llist :: "'a item set => 'a item set"
- LListD :: "('a item * 'a item)set => ('a item * 'a item)set"
-
+ for A :: "'a item set"
+ where
+ NIL_I: "NIL \<in> llist(A)"
+ | CONS_I: "[| a \<in> A; M \<in> llist(A) |] ==> CONS a M \<in> llist(A)"
-coinductive "llist(A)"
- intros
- NIL_I: "NIL \<in> llist(A)"
- CONS_I: "[| a \<in> A; M \<in> llist(A) |] ==> CONS a M \<in> llist(A)"
-
-coinductive "LListD(r)"
- intros
+coinductive_set
+ LListD :: "('a item * 'a item)set => ('a item * 'a item)set"
+ for r :: "('a item * 'a item)set"
+ where
NIL_I: "(NIL, NIL) \<in> LListD(r)"
- CONS_I: "[| (a,b) \<in> r; (M,N) \<in> LListD(r) |]
+ | CONS_I: "[| (a,b) \<in> r; (M,N) \<in> LListD(r) |]
==> (CONS a M, CONS b N) \<in> LListD(r)"
@@ -159,11 +157,19 @@
declare option.split [split]
text{*This justifies using llist in other recursive type definitions*}
-lemma llist_mono: "A\<subseteq>B ==> llist(A) \<subseteq> llist(B)"
-apply (simp add: llist.defs)
-apply (rule gfp_mono)
-apply (assumption | rule basic_monos)+
-done
+lemma llist_mono:
+ assumes subset: "A \<subseteq> B"
+ shows "llist A \<subseteq> llist B"
+proof
+ fix x
+ assume "x \<in> llist A"
+ then show "x \<in> llist B"
+ proof coinduct
+ case llist
+ then show ?case using subset
+ by cases blast+
+ qed
+qed
lemma llist_unfold: "llist(A) = usum {Numb(0)} (uprod A (llist A))"
@@ -195,9 +201,9 @@
text{*Utilise the ``strong'' part, i.e. @{text "gfp(f)"}*}
lemma list_Fun_llist_I: "M \<in> llist(A) ==> M \<in> list_Fun A (X Un llist(A))"
-apply (unfold llist.defs list_Fun_def)
-apply (rule gfp_fun_UnI2)
-apply (rule monoI, auto)
+apply (unfold list_Fun_def)
+apply (erule llist.cases)
+apply auto
done
subsection{* @{text LList_corec} satisfies the desired recurion equation *}
@@ -278,10 +284,10 @@
text{*The domain of the @{text LListD} relation*}
lemma Domain_LListD:
"Domain (LListD(diag A)) \<subseteq> llist(A)"
-apply (simp add: llist.defs NIL_def CONS_def)
-apply (rule gfp_upperbound)
-txt{*avoids unfolding @{text LListD} on the rhs*}
-apply (rule_tac P = "%x. Domain x \<subseteq> ?B" in LListD_unfold [THEN ssubst], auto)
+apply (rule subsetI)
+apply (erule llist.coinduct)
+apply (simp add: NIL_def CONS_def)
+apply (drule_tac P = "%x. xa \<in> Domain x" in LListD_unfold [THEN subst], auto)
done
text{*This inclusion justifies the use of coinduction to show @{text "M = N"}*}
@@ -305,6 +311,7 @@
lemma LListD_coinduct:
"[| M \<in> X; X \<subseteq> LListD_Fun r (X Un LListD(r)) |] ==> M \<in> LListD(r)"
+apply (cases M)
apply (simp add: LListD_Fun_def)
apply (erule LListD.coinduct)
apply (auto );
@@ -320,9 +327,10 @@
text{*Utilise the "strong" part, i.e. @{text "gfp(f)"}*}
lemma LListD_Fun_LListD_I:
"M \<in> LListD(r) ==> M \<in> LListD_Fun r (X Un LListD(r))"
-apply (unfold LListD.defs LListD_Fun_def)
-apply (rule gfp_fun_UnI2)
-apply (rule monoI, auto)
+apply (cases M)
+apply (simp add: LListD_Fun_def)
+apply (erule LListD.cases)
+apply auto
done
@@ -523,7 +531,7 @@
apply (rule LList_equalityI)
apply (erule imageI)
apply (rule image_subsetI)
-apply (erule_tac aa=x in llist.cases)
+apply (erule_tac a=x in llist.cases)
apply (erule ssubst, erule ssubst, erule LListD_Fun_diag_I, blast)
done
@@ -608,8 +616,8 @@
apply (rule_tac X = "\<Union>u\<in>llist (A) . \<Union>v \<in> llist (A) . {Lappend u v}" in llist_coinduct)
apply fast
apply safe
-apply (erule_tac aa = u in llist.cases)
-apply (erule_tac aa = v in llist.cases, simp_all, blast)
+apply (erule_tac a = u in llist.cases)
+apply (erule_tac a = v in llist.cases, simp_all, blast)
done
text{*strong co-induction: bisimulation and case analysis on one variable*}
@@ -617,7 +625,7 @@
apply (rule_tac X = "(%u. Lappend u N) `llist (A)" in llist_coinduct)
apply (erule imageI)
apply (rule image_subsetI)
-apply (erule_tac aa = x in llist.cases)
+apply (erule_tac a = x in llist.cases)
apply (simp add: list_Fun_llist_I, simp)
done
--- a/src/HOL/Induct/Mutil.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/Mutil.thy Wed Jul 11 11:14:51 2007 +0200
@@ -15,18 +15,19 @@
the Mutilated Checkerboard Problem by J McCarthy.
*}
-consts tiling :: "'a set set => 'a set set"
-inductive "tiling A"
- intros
+inductive_set
+ tiling :: "'a set set => 'a set set"
+ for A :: "'a set set"
+ where
empty [simp, intro]: "{} \<in> tiling A"
- Un [simp, intro]: "[| a \<in> A; t \<in> tiling A; a \<inter> t = {} |]
+ | Un [simp, intro]: "[| a \<in> A; t \<in> tiling A; a \<inter> t = {} |]
==> a \<union> t \<in> tiling A"
-consts domino :: "(nat \<times> nat) set set"
-inductive domino
- intros
+inductive_set
+ domino :: "(nat \<times> nat) set set"
+ where
horiz [simp]: "{(i, j), (i, Suc j)} \<in> domino"
- vertl [simp]: "{(i, j), (Suc i, j)} \<in> domino"
+ | vertl [simp]: "{(i, j), (Suc i, j)} \<in> domino"
text {* \medskip Sets of squares of the given colour*}
--- a/src/HOL/Induct/PropLog.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/PropLog.thy Wed Jul 11 11:14:51 2007 +0200
@@ -26,20 +26,15 @@
subsection {* The proof system *}
-consts
- thms :: "'a pl set => 'a pl set"
-
-abbreviation
- thm_rel :: "['a pl set, 'a pl] => bool" (infixl "|-" 50) where
- "H |- p == p \<in> thms H"
-
-inductive "thms(H)"
- intros
- H [intro]: "p\<in>H ==> H |- p"
- K: "H |- p->q->p"
- S: "H |- (p->q->r) -> (p->q) -> p->r"
- DN: "H |- ((p->false) -> false) -> p"
- MP: "[| H |- p->q; H |- p |] ==> H |- q"
+inductive
+ thms :: "['a pl set, 'a pl] => bool" (infixl "|-" 50)
+ for H :: "'a pl set"
+ where
+ H [intro]: "p\<in>H ==> H |- p"
+ | K: "H |- p->q->p"
+ | S: "H |- (p->q->r) -> (p->q) -> p->r"
+ | DN: "H |- ((p->false) -> false) -> p"
+ | MP: "[| H |- p->q; H |- p |] ==> H |- q"
subsection {* The semantics *}
@@ -80,9 +75,9 @@
subsection {* Proof theory of propositional logic *}
lemma thms_mono: "G<=H ==> thms(G) <= thms(H)"
-apply (unfold thms.defs )
-apply (rule lfp_mono)
-apply (assumption | rule basic_monos)+
+apply (rule predicate1I)
+apply (erule thms.induct)
+apply (auto intro: thms.intros)
done
lemma thms_I: "H |- p->p"
@@ -94,7 +89,7 @@
lemma weaken_left: "[| G \<subseteq> H; G|-p |] ==> H|-p"
-- {* Order of premises is convenient with @{text THEN} *}
- by (erule thms_mono [THEN subsetD])
+ by (erule thms_mono [THEN predicate1D])
lemmas weaken_left_insert = subset_insertI [THEN weaken_left]
--- a/src/HOL/Induct/QuoDataType.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/QuoDataType.thy Wed Jul 11 11:14:51 2007 +0200
@@ -19,31 +19,25 @@
| DECRYPT nat freemsg
text{*The equivalence relation, which makes encryption and decryption inverses
-provided the keys are the same.*}
-consts msgrel :: "(freemsg * freemsg) set"
-
-abbreviation
- msg_rel :: "[freemsg, freemsg] => bool" (infixl "~~" 50) where
- "X ~~ Y == (X,Y) \<in> msgrel"
+provided the keys are the same.
-notation (xsymbols)
- msg_rel (infixl "\<sim>" 50)
-notation (HTML output)
- msg_rel (infixl "\<sim>" 50)
-
-text{*The first two rules are the desired equations. The next four rules
+The first two rules are the desired equations. The next four rules
make the equations applicable to subterms. The last two rules are symmetry
and transitivity.*}
-inductive "msgrel"
- intros
- CD: "CRYPT K (DECRYPT K X) \<sim> X"
- DC: "DECRYPT K (CRYPT K X) \<sim> X"
- NONCE: "NONCE N \<sim> NONCE N"
- MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
- CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
- DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
- SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X"
- TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
+
+inductive_set
+ msgrel :: "(freemsg * freemsg) set"
+ and msg_rel :: "[freemsg, freemsg] => bool" (infixl "\<sim>" 50)
+ where
+ "X \<sim> Y == (X,Y) \<in> msgrel"
+ | CD: "CRYPT K (DECRYPT K X) \<sim> X"
+ | DC: "DECRYPT K (CRYPT K X) \<sim> X"
+ | NONCE: "NONCE N \<sim> NONCE N"
+ | MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
+ | CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
+ | DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
+ | SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X"
+ | TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
text{*Proving that it is an equivalence relation*}
--- a/src/HOL/Induct/QuoNestedDataType.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/QuoNestedDataType.thy Wed Jul 11 11:14:51 2007 +0200
@@ -18,28 +18,21 @@
| FNCALL nat "freeExp list"
text{*The equivalence relation, which makes PLUS associative.*}
-consts exprel :: "(freeExp * freeExp) set"
-
-abbreviation
- exp_rel :: "[freeExp, freeExp] => bool" (infixl "~~" 50) where
- "X ~~ Y == (X,Y) \<in> exprel"
-
-notation (xsymbols)
- exp_rel (infixl "\<sim>" 50)
-notation (HTML output)
- exp_rel (infixl "\<sim>" 50)
text{*The first rule is the desired equation. The next three rules
make the equations applicable to subterms. The last two rules are symmetry
and transitivity.*}
-inductive "exprel"
- intros
- ASSOC: "PLUS X (PLUS Y Z) \<sim> PLUS (PLUS X Y) Z"
- VAR: "VAR N \<sim> VAR N"
- PLUS: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> PLUS X Y \<sim> PLUS X' Y'"
- FNCALL: "(Xs,Xs') \<in> listrel exprel \<Longrightarrow> FNCALL F Xs \<sim> FNCALL F Xs'"
- SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X"
- TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
+inductive_set
+ exprel :: "(freeExp * freeExp) set"
+ and exp_rel :: "[freeExp, freeExp] => bool" (infixl "\<sim>" 50)
+ where
+ "X \<sim> Y == (X,Y) \<in> exprel"
+ | ASSOC: "PLUS X (PLUS Y Z) \<sim> PLUS (PLUS X Y) Z"
+ | VAR: "VAR N \<sim> VAR N"
+ | PLUS: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> PLUS X Y \<sim> PLUS X' Y'"
+ | FNCALL: "(Xs,Xs') \<in> listrel exprel \<Longrightarrow> FNCALL F Xs \<sim> FNCALL F Xs'"
+ | SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X"
+ | TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
monos listrel_mono
@@ -47,7 +40,7 @@
lemma exprel_refl: "X \<sim> X"
and list_exprel_refl: "(Xs,Xs) \<in> listrel(exprel)"
- by (induct X and Xs) (blast intro: exprel.intros listrel_intros)+
+ by (induct X and Xs) (blast intro: exprel.intros listrel.intros)+
theorem equiv_exprel: "equiv UNIV exprel"
proof -
@@ -63,13 +56,13 @@
lemma FNCALL_Nil: "FNCALL F [] \<sim> FNCALL F []"
apply (rule exprel.intros)
-apply (rule listrel_intros)
+apply (rule listrel.intros)
done
lemma FNCALL_Cons:
"\<lbrakk>X \<sim> X'; (Xs,Xs') \<in> listrel(exprel)\<rbrakk>
\<Longrightarrow> FNCALL F (X#Xs) \<sim> FNCALL F (X'#Xs')"
-by (blast intro: exprel.intros listrel_intros)
+by (blast intro: exprel.intros listrel.intros)
@@ -98,7 +91,7 @@
(the abstract constructor) is injective*}
theorem exprel_imp_eq_freevars: "U \<sim> V \<Longrightarrow> freevars U = freevars V"
apply (induct set: exprel)
-apply (erule_tac [4] listrel_induct)
+apply (erule_tac [4] listrel.induct)
apply (simp_all add: Un_assoc)
done
@@ -129,7 +122,7 @@
theorem exprel_imp_eq_freefun:
"U \<sim> V \<Longrightarrow> freefun U = freefun V"
- by (induct set: exprel) (simp_all add: listrel_intros)
+ by (induct set: exprel) (simp_all add: listrel.intros)
text{*This function, which returns the list of function arguments, is used to
@@ -143,8 +136,8 @@
theorem exprel_imp_eqv_freeargs:
"U \<sim> V \<Longrightarrow> (freeargs U, freeargs V) \<in> listrel exprel"
apply (induct set: exprel)
-apply (erule_tac [4] listrel_induct)
-apply (simp_all add: listrel_intros)
+apply (erule_tac [4] listrel.induct)
+apply (simp_all add: listrel.intros)
apply (blast intro: symD [OF equiv.sym [OF equiv_list_exprel]])
apply (blast intro: transD [OF equiv.trans [OF equiv_list_exprel]])
done
@@ -258,7 +251,7 @@
"FnCall F [Abs_Exp(exprel``{U})] = Abs_Exp (exprel``{FNCALL F [U]})"
proof -
have "(\<lambda>U. exprel `` {FNCALL F [U]}) respects exprel"
- by (simp add: congruent_def FNCALL_Cons listrel_intros)
+ by (simp add: congruent_def FNCALL_Cons listrel.intros)
thus ?thesis
by (simp add: FnCall_def UN_equiv_class [OF equiv_exprel])
qed
--- a/src/HOL/Induct/ROOT.ML Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/ROOT.ML Wed Jul 11 11:14:51 2007 +0200
@@ -13,3 +13,4 @@
time_use_thy "PropLog";
time_use_thy "SList";
time_use_thy "LFilter";
+time_use_thy "Com";
--- a/src/HOL/Induct/SList.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/SList.thy Wed Jul 11 11:14:51 2007 +0200
@@ -46,12 +46,12 @@
CONS :: "['a item, 'a item] => 'a item" where
"CONS M N = In1(Scons M N)"
-consts
- list :: "'a item set => 'a item set"
-inductive "list(A)"
- intros
+inductive_set
+ 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"
+ | CONS_I: "[| a: A; M: list A |] ==> CONS a M : list A"
typedef (List)
@@ -149,6 +149,8 @@
ttl :: "'a list => 'a list" where
"ttl xs = list_rec xs [] (%x xs r. xs)"
+(*<*)no_syntax
+ member :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infixl "mem" 55)(*>*)
definition
member :: "['a, 'a list] => bool" (infixl "mem" 55) where
"x mem xs = list_rec xs False (%y ys r. if y=x then True else r)"
@@ -254,16 +256,17 @@
(*This justifies using list in other recursive type definitions*)
lemma list_mono: "A<=B ==> list(A) <= list(B)"
-apply (unfold list.defs )
-apply (rule lfp_mono)
-apply (assumption | rule basic_monos)+
+apply (rule subsetI)
+apply (erule list.induct)
+apply (auto intro!: list.intros)
done
(*Type checking -- list creates well-founded sets*)
lemma list_sexp: "list(sexp) <= sexp"
-apply (unfold NIL_def CONS_def list.defs)
-apply (rule lfp_lowerbound)
-apply (fast intro: sexp.intros sexp_In0I sexp_In1I)
+apply (rule subsetI)
+apply (erule list.induct)
+apply (unfold NIL_def CONS_def)
+apply (auto intro: sexp.intros sexp_In0I sexp_In1I)
done
(* A <= sexp ==> list(A) <= sexp *)
--- a/src/HOL/Induct/Sexp.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/Sexp.thy Wed Jul 11 11:14:51 2007 +0200
@@ -14,14 +14,12 @@
abbreviation "Leaf == Datatype.Leaf"
abbreviation "Numb == Datatype.Numb"
-consts
+inductive_set
sexp :: "'a item set"
-
-inductive sexp
- intros
+ where
LeafI: "Leaf(a) \<in> sexp"
- NumbI: "Numb(i) \<in> sexp"
- SconsI: "[| M \<in> sexp; N \<in> sexp |] ==> Scons M N \<in> sexp"
+ | NumbI: "Numb(i) \<in> sexp"
+ | SconsI: "[| M \<in> sexp; N \<in> sexp |] ==> Scons M N \<in> sexp"
definition
sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b,
--- a/src/HOL/Induct/Sigma_Algebra.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Induct/Sigma_Algebra.thy Wed Jul 11 11:14:51 2007 +0200
@@ -13,15 +13,14 @@
\<sigma>}-algebra over a given set of sets.
*}
-consts
+inductive_set
\<sigma>_algebra :: "'a set set => 'a set set"
-
-inductive "\<sigma>_algebra A"
- intros
+ for A :: "'a set set"
+ where
basic: "a \<in> A ==> a \<in> \<sigma>_algebra A"
- UNIV: "UNIV \<in> \<sigma>_algebra A"
- complement: "a \<in> \<sigma>_algebra A ==> -a \<in> \<sigma>_algebra A"
- Union: "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Union>i. a i) \<in> \<sigma>_algebra A"
+ | UNIV: "UNIV \<in> \<sigma>_algebra A"
+ | complement: "a \<in> \<sigma>_algebra A ==> -a \<in> \<sigma>_algebra A"
+ | Union: "(!!i::nat. a i \<in> \<sigma>_algebra A) ==> (\<Union>i. a i) \<in> \<sigma>_algebra A"
text {*
The following basic facts are consequences of the closure properties
--- a/src/HOL/Isar_examples/MutilatedCheckerboard.thy Wed Jul 11 11:13:08 2007 +0200
+++ b/src/HOL/Isar_examples/MutilatedCheckerboard.thy Wed Jul 11 11:14:51 2007 +0200
@@ -17,13 +17,12 @@
subsection {* Tilings *}
-consts
+inductive_set
tiling :: "'a set set => 'a set set"
-
-inductive "tiling A"
- intros
+ for A :: "'a set set"
+ where
empty: "{} : tiling A"
- Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
+ | Un: "a : A ==> t : tiling A ==> a <= - t ==> a Un t : tiling A"
text "The union of two disjoint tilings is a tiling."
@@ -118,13 +117,11 @@
subsection {* Dominoes *}
-consts
+inductive_set
domino :: "(nat * nat) set set"
-
-inductive domino
- intros
+ where
horiz: "{(i, j), (i, j + 1)} : domino"
- vertl: "{(i, j), (i + 1, j)} : domino"
+ | vertl: "{(i, j), (i + 1, j)} : domino"
lemma dominoes_tile_row:
"{i} <*> below (2 * n) : tiling domino"