# HG changeset patch # User huffman # Date 1268546434 28800 # Node ID 38538bfe9ca600b5a2b6bf6f5b50d1c798320380 # Parent 8a314dd86714bd213d96d9e87b517798e4db6bbf declare case_names for various induction rules diff -r 8a314dd86714 -r 38538bfe9ca6 src/HOLCF/One.thy --- a/src/HOLCF/One.thy Sat Mar 13 21:07:20 2010 -0800 +++ b/src/HOLCF/One.thy Sat Mar 13 22:00:34 2010 -0800 @@ -22,10 +22,10 @@ lemma Exh_one: "t = \ \ t = ONE" unfolding ONE_def by (induct t) simp_all -lemma oneE: "\p = \ \ Q; p = ONE \ Q\ \ Q" +lemma oneE [case_names bottom ONE]: "\p = \ \ Q; p = ONE \ Q\ \ Q" unfolding ONE_def by (induct p) simp_all -lemma one_induct: "\P \; P ONE\ \ P x" +lemma one_induct [case_names bottom ONE]: "\P \; P ONE\ \ P x" by (cases x rule: oneE) simp_all lemma dist_below_one [simp]: "\ ONE \ \" diff -r 8a314dd86714 -r 38538bfe9ca6 src/HOLCF/Sprod.thy --- a/src/HOLCF/Sprod.thy Sat Mar 13 21:07:20 2010 -0800 +++ b/src/HOLCF/Sprod.thy Sat Mar 13 22:00:34 2010 -0800 @@ -80,11 +80,11 @@ apply fast done -lemma sprodE [cases type: sprod]: +lemma sprodE [case_names bottom spair, cases type: sprod]: "\p = \ \ Q; \x y. \p = (:x, y:); x \ \; y \ \\ \ Q\ \ Q" -by (cut_tac z=p in Exh_Sprod, auto) +using Exh_Sprod [of p] by auto -lemma sprod_induct [induct type: sprod]: +lemma sprod_induct [case_names bottom spair, induct type: sprod]: "\P \; \x y. \x \ \; y \ \\ \ P (:x, y:)\ \ P x" by (cases x, simp_all) diff -r 8a314dd86714 -r 38538bfe9ca6 src/HOLCF/Ssum.thy --- a/src/HOLCF/Ssum.thy Sat Mar 13 21:07:20 2010 -0800 +++ b/src/HOLCF/Ssum.thy Sat Mar 13 22:00:34 2010 -0800 @@ -151,19 +151,19 @@ apply (simp add: sinr_Abs_Ssum Ssum_def) done -lemma ssumE [cases type: ssum]: +lemma ssumE [case_names bottom sinl sinr, cases type: ssum]: "\p = \ \ Q; \x. \p = sinl\x; x \ \\ \ Q; \y. \p = sinr\y; y \ \\ \ Q\ \ Q" -by (cut_tac z=p in Exh_Ssum, auto) +using Exh_Ssum [of p] by auto -lemma ssum_induct [induct type: ssum]: +lemma ssum_induct [case_names bottom sinl sinr, induct type: ssum]: "\P \; \x. x \ \ \ P (sinl\x); \y. y \ \ \ P (sinr\y)\ \ P x" by (cases x, simp_all) -lemma ssumE2: +lemma ssumE2 [case_names sinl sinr]: "\\x. p = sinl\x \ Q; \y. p = sinr\y \ Q\ \ Q" by (cases p, simp only: sinl_strict [symmetric], simp, simp) diff -r 8a314dd86714 -r 38538bfe9ca6 src/HOLCF/Tr.thy --- a/src/HOLCF/Tr.thy Sat Mar 13 21:07:20 2010 -0800 +++ b/src/HOLCF/Tr.thy Sat Mar 13 22:00:34 2010 -0800 @@ -29,10 +29,12 @@ lemma Exh_tr: "t = \ \ t = TT \ t = FF" unfolding FF_def TT_def by (induct t) auto -lemma trE: "\p = \ \ Q; p = TT \ Q; p = FF \ Q\ \ Q" +lemma trE [case_names bottom TT FF]: + "\p = \ \ Q; p = TT \ Q; p = FF \ Q\ \ Q" unfolding FF_def TT_def by (induct p) auto -lemma tr_induct: "\P \; P TT; P FF\ \ P x" +lemma tr_induct [case_names bottom TT FF]: + "\P \; P TT; P FF\ \ P x" by (cases x rule: trE) simp_all text {* distinctness for type @{typ tr} *} diff -r 8a314dd86714 -r 38538bfe9ca6 src/HOLCF/Up.thy --- a/src/HOLCF/Up.thy Sat Mar 13 21:07:20 2010 -0800 +++ b/src/HOLCF/Up.thy Sat Mar 13 22:00:34 2010 -0800 @@ -237,13 +237,15 @@ lemma up_below [simp]: "up\x \ up\y \ x \ y" by (simp add: up_def cont_Iup) -lemma upE [cases type: u]: "\p = \ \ Q; \x. p = up\x \ Q\ \ Q" +lemma upE [case_names bottom up, cases type: u]: + "\p = \ \ Q; \x. p = up\x \ Q\ \ Q" apply (cases p) apply (simp add: inst_up_pcpo) apply (simp add: up_def cont_Iup) done -lemma up_induct [induct type: u]: "\P \; \x. P (up\x)\ \ P x" +lemma up_induct [case_names bottom up, induct type: u]: + "\P \; \x. P (up\x)\ \ P x" by (cases x, simp_all) text {* lifting preserves chain-finiteness *}