# HG changeset patch # User wenzelm # Date 1181752217 -7200 # Node ID 53317a1ec8b2d456049f3d27c2afdf5c1c8b2972 # Parent 45cd7db985b38ca2d698436b901f0b0172a0f214 tuned proofs: avoid implicit prems; tuned; diff -r 45cd7db985b3 -r 53317a1ec8b2 src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Wed Jun 13 18:30:16 2007 +0200 +++ b/src/HOLCF/Domain.thy Wed Jun 13 18:30:17 2007 +0200 @@ -11,6 +11,7 @@ defaultsort pcpo + subsection {* Continuous isomorphisms *} text {* A locale for continuous isomorphisms *} @@ -20,41 +21,42 @@ fixes rep :: "'b \ 'a" assumes abs_iso [simp]: "rep\(abs\x) = x" assumes rep_iso [simp]: "abs\(rep\y) = y" +begin -lemma (in iso) swap: "iso rep abs" -by (rule iso.intro [OF rep_iso abs_iso]) +lemma swap: "iso rep abs" + by (rule iso.intro [OF rep_iso abs_iso]) -lemma (in iso) abs_less: "(abs\x \ abs\y) = (x \ y)" +lemma abs_less: "(abs\x \ abs\y) = (x \ y)" proof assume "abs\x \ abs\y" - hence "rep\(abs\x) \ rep\(abs\y)" by (rule monofun_cfun_arg) - thus "x \ y" by simp + then have "rep\(abs\x) \ rep\(abs\y)" by (rule monofun_cfun_arg) + then show "x \ y" by simp next assume "x \ y" - thus "abs\x \ abs\y" by (rule monofun_cfun_arg) + then show "abs\x \ abs\y" by (rule monofun_cfun_arg) qed -lemma (in iso) rep_less: "(rep\x \ rep\y) = (x \ y)" -by (rule iso.abs_less [OF swap]) +lemma rep_less: "(rep\x \ rep\y) = (x \ y)" + by (rule iso.abs_less [OF swap]) -lemma (in iso) abs_eq: "(abs\x = abs\y) = (x = y)" -by (simp add: po_eq_conv abs_less) +lemma abs_eq: "(abs\x = abs\y) = (x = y)" + by (simp add: po_eq_conv abs_less) -lemma (in iso) rep_eq: "(rep\x = rep\y) = (x = y)" -by (rule iso.abs_eq [OF swap]) +lemma rep_eq: "(rep\x = rep\y) = (x = y)" + by (rule iso.abs_eq [OF swap]) -lemma (in iso) abs_strict: "abs\\ = \" +lemma abs_strict: "abs\\ = \" proof - have "\ \ rep\\" .. - hence "abs\\ \ abs\(rep\\)" by (rule monofun_cfun_arg) - hence "abs\\ \ \" by simp - thus ?thesis by (rule UU_I) + then have "abs\\ \ abs\(rep\\)" by (rule monofun_cfun_arg) + then have "abs\\ \ \" by simp + then show ?thesis by (rule UU_I) qed -lemma (in iso) rep_strict: "rep\\ = \" -by (rule iso.abs_strict [OF swap]) +lemma rep_strict: "rep\\ = \" + by (rule iso.abs_strict [OF swap]) -lemma (in iso) abs_defin': "abs\x = \ \ x = \" +lemma abs_defin': "abs\x = \ \ x = \" proof - have "x = rep\(abs\x)" by simp also assume "abs\x = \" @@ -62,49 +64,52 @@ finally show "x = \" . qed -lemma (in iso) rep_defin': "rep\z = \ \ z = \" -by (rule iso.abs_defin' [OF swap]) +lemma rep_defin': "rep\z = \ \ z = \" + by (rule iso.abs_defin' [OF swap]) -lemma (in iso) abs_defined: "z \ \ \ abs\z \ \" -by (erule contrapos_nn, erule abs_defin') +lemma abs_defined: "z \ \ \ abs\z \ \" + by (erule contrapos_nn, erule abs_defin') -lemma (in iso) rep_defined: "z \ \ \ rep\z \ \" -by (rule iso.abs_defined [OF iso.swap]) +lemma rep_defined: "z \ \ \ rep\z \ \" + by (rule iso.abs_defined [OF iso.swap]) (rule iso_axioms) -lemma (in iso) abs_defined_iff: "(abs\x = \) = (x = \)" -by (auto elim: abs_defin' intro: abs_strict) +lemma abs_defined_iff: "(abs\x = \) = (x = \)" + by (auto elim: abs_defin' intro: abs_strict) -lemma (in iso) rep_defined_iff: "(rep\x = \) = (x = \)" -by (rule iso.abs_defined_iff [OF iso.swap]) +lemma rep_defined_iff: "(rep\x = \) = (x = \)" + by (rule iso.abs_defined_iff [OF iso.swap]) (rule iso_axioms) lemma (in iso) compact_abs_rev: "compact (abs\x) \ compact x" proof (unfold compact_def) assume "adm (\y. \ abs\x \ y)" with cont_Rep_CFun2 have "adm (\y. \ abs\x \ abs\y)" by (rule adm_subst) - thus "adm (\y. \ x \ y)" using abs_less by simp + then show "adm (\y. \ x \ y)" using abs_less by simp qed -lemma (in iso) compact_rep_rev: "compact (rep\x) \ compact x" -by (rule iso.compact_abs_rev [OF iso.swap]) +lemma compact_rep_rev: "compact (rep\x) \ compact x" + by (rule iso.compact_abs_rev [OF iso.swap]) (rule iso_axioms) -lemma (in iso) compact_abs: "compact x \ compact (abs\x)" -by (rule compact_rep_rev, simp) +lemma compact_abs: "compact x \ compact (abs\x)" + by (rule compact_rep_rev) simp -lemma (in iso) compact_rep: "compact x \ compact (rep\x)" -by (rule iso.compact_abs [OF iso.swap]) +lemma compact_rep: "compact x \ compact (rep\x)" + by (rule iso.compact_abs [OF iso.swap]) (rule iso_axioms) -lemma (in iso) iso_swap: "(x = abs\y) = (rep\x = y)" +lemma iso_swap: "(x = abs\y) = (rep\x = y)" proof assume "x = abs\y" - hence "rep\x = rep\(abs\y)" by simp - thus "rep\x = y" by simp + then have "rep\x = rep\(abs\y)" by simp + then show "rep\x = y" by simp next assume "rep\x = y" - hence "abs\(rep\x) = abs\y" by simp - thus "x = abs\y" by simp + then have "abs\(rep\x) = abs\y" by simp + then show "x = abs\y" by simp qed +end + + subsection {* Casedist *} lemma ex_one_defined_iff: @@ -114,7 +119,7 @@ apply simp apply simp apply force -done + done lemma ex_up_defined_iff: "(\x. P x \ x \ \) = (\x. P (up\x))" @@ -123,7 +128,7 @@ apply simp apply fast apply (force intro!: up_defined) -done + done lemma ex_sprod_defined_iff: "(\y. P y \ y \ \) = @@ -133,7 +138,7 @@ apply simp apply fast apply (force intro!: spair_defined) -done + done lemma ex_sprod_up_defined_iff: "(\y. P y \ y \ \) = @@ -145,7 +150,7 @@ apply simp apply fast apply (force intro!: spair_defined) -done + done lemma ex_ssum_defined_iff: "(\x. P x \ x \ \) = @@ -161,10 +166,10 @@ apply (erule disjE) apply force apply force -done + done lemma exh_start: "p = \ \ (\x. p = x \ x \ \)" -by auto + by auto lemmas ex_defined_iffs = ex_ssum_defined_iff @@ -176,16 +181,16 @@ text {* Rules for turning exh into casedist *} lemma exh_casedist0: "\R; R \ P\ \ P" (* like make_elim *) -by auto + by auto lemma exh_casedist1: "((P \ Q \ R) \ S) \ (\P \ R; Q \ R\ \ S)" -by rule auto + by rule auto lemma exh_casedist2: "(\x. P x \ Q) \ (\x. P x \ Q)" -by rule auto + by rule auto lemma exh_casedist3: "(P \ Q \ R) \ (P \ Q \ R)" -by rule auto + by rule auto lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3