New file for theorems used by the domain package
authorhuffman
Sat, 16 Apr 2005 00:16:44 +0200
changeset 1574129a78517543f
parent 15740 d63e7a65b2d0
child 15742 64eae3513064
New file for theorems used by the domain package
src/HOLCF/Domain.thy
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOLCF/Domain.thy	Sat Apr 16 00:16:44 2005 +0200
     1.3 @@ -0,0 +1,180 @@
     1.4 +(*  Title:      HOLCF/Domain.thy
     1.5 +    ID:         $Id$
     1.6 +    Author:     Brian Huffman
     1.7 +    License:    GPL (GNU GENERAL PUBLIC LICENSE)
     1.8 +*)
     1.9 +
    1.10 +header {* Domain package *}
    1.11 +
    1.12 +theory Domain
    1.13 +imports Ssum Sprod One Up
    1.14 +files
    1.15 +  ("domain/library.ML")
    1.16 +  ("domain/syntax.ML")
    1.17 +  ("domain/axioms.ML")
    1.18 +  ("domain/theorems.ML")
    1.19 +  ("domain/extender.ML")
    1.20 +  ("domain/interface.ML")
    1.21 +begin
    1.22 +
    1.23 +defaultsort pcpo
    1.24 +
    1.25 +subsection {* Continuous isomorphisms *}
    1.26 +
    1.27 +text {* A locale for continuous isomorphisms *}
    1.28 +
    1.29 +locale iso =
    1.30 +  fixes abs :: "'a \<rightarrow> 'b"
    1.31 +  fixes rep :: "'b \<rightarrow> 'a"
    1.32 +  assumes abs_iso [simp]: "rep\<cdot>(abs\<cdot>x) = x"
    1.33 +  assumes rep_iso [simp]: "abs\<cdot>(rep\<cdot>y) = y"
    1.34 +
    1.35 +lemma (in iso) swap: "iso rep abs"
    1.36 +by (rule iso.intro [OF rep_iso abs_iso])
    1.37 +
    1.38 +lemma (in iso) abs_strict: "abs\<cdot>\<bottom> = \<bottom>"
    1.39 +proof -
    1.40 +  have "\<bottom> \<sqsubseteq> rep\<cdot>\<bottom>" ..
    1.41 +  hence "abs\<cdot>\<bottom> \<sqsubseteq> abs\<cdot>(rep\<cdot>\<bottom>)" by (rule monofun_cfun_arg)
    1.42 +  hence "abs\<cdot>\<bottom> \<sqsubseteq> \<bottom>" by simp
    1.43 +  thus ?thesis by (rule UU_I)
    1.44 +qed
    1.45 +
    1.46 +lemma (in iso) rep_strict: "rep\<cdot>\<bottom> = \<bottom>"
    1.47 +by (rule iso.abs_strict [OF swap])
    1.48 +
    1.49 +lemma (in iso) abs_defin': "abs\<cdot>z = \<bottom> \<Longrightarrow> z = \<bottom>"
    1.50 +proof -
    1.51 +  assume A: "abs\<cdot>z = \<bottom>"
    1.52 +  have "z = rep\<cdot>(abs\<cdot>z)" by simp
    1.53 +  also have "\<dots> = rep\<cdot>\<bottom>" by (simp only: A)
    1.54 +  also note rep_strict
    1.55 +  finally show "z = \<bottom>" .
    1.56 +qed
    1.57 +
    1.58 +lemma (in iso) rep_defin': "rep\<cdot>z = \<bottom> \<Longrightarrow> z = \<bottom>"
    1.59 +by (rule iso.abs_defin' [OF swap])
    1.60 +
    1.61 +lemma (in iso) abs_defined: "z \<noteq> \<bottom> \<Longrightarrow> abs\<cdot>z \<noteq> \<bottom>"
    1.62 +by (erule contrapos_nn, erule abs_defin')
    1.63 +
    1.64 +lemma (in iso) rep_defined: "z \<noteq> \<bottom> \<Longrightarrow> rep\<cdot>z \<noteq> \<bottom>"
    1.65 +by (erule contrapos_nn, erule rep_defin')
    1.66 +
    1.67 +lemma (in iso) iso_swap: "(x = abs\<cdot>y) = (rep\<cdot>x = y)"
    1.68 +proof
    1.69 +  assume "x = abs\<cdot>y"
    1.70 +  hence "rep\<cdot>x = rep\<cdot>(abs\<cdot>y)" by simp
    1.71 +  thus "rep\<cdot>x = y" by simp
    1.72 +next
    1.73 +  assume "rep\<cdot>x = y"
    1.74 +  hence "abs\<cdot>(rep\<cdot>x) = abs\<cdot>y" by simp
    1.75 +  thus "x = abs\<cdot>y" by simp
    1.76 +qed
    1.77 +
    1.78 +subsection {* Casedist *}
    1.79 +
    1.80 +lemma ex_one_defined_iff:
    1.81 +  "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = P ONE"
    1.82 + apply safe
    1.83 +  apply (rule_tac p=x in oneE)
    1.84 +   apply simp
    1.85 +  apply simp
    1.86 + apply force
    1.87 +done
    1.88 +
    1.89 +lemma ex_up_defined_iff:
    1.90 +  "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = (\<exists>x. P (up\<cdot>x))"
    1.91 + apply safe
    1.92 +  apply (rule_tac p=x in upE1)
    1.93 +   apply simp
    1.94 +  apply fast
    1.95 + apply (force intro!: defined_up)
    1.96 +done
    1.97 +
    1.98 +lemma ex_sprod_defined_iff:
    1.99 + "(\<exists>y. P y \<and> y \<noteq> \<bottom>) =
   1.100 +  (\<exists>x y. (P (:x, y:) \<and> x \<noteq> \<bottom>) \<and> y \<noteq> \<bottom>)"
   1.101 + apply safe
   1.102 +  apply (rule_tac p=y in sprodE)
   1.103 +   apply simp
   1.104 +  apply fast
   1.105 + apply (force intro!: defined_spair)
   1.106 +done
   1.107 +
   1.108 +lemma ex_sprod_up_defined_iff:
   1.109 + "(\<exists>y. P y \<and> y \<noteq> \<bottom>) =
   1.110 +  (\<exists>x y. P (:up\<cdot>x, y:) \<and> y \<noteq> \<bottom>)"
   1.111 + apply safe
   1.112 +  apply (rule_tac p=y in sprodE)
   1.113 +   apply simp
   1.114 +  apply (rule_tac p=x in upE1)
   1.115 +   apply simp
   1.116 +  apply fast
   1.117 + apply (force intro!: defined_spair)
   1.118 +done
   1.119 +
   1.120 +lemma ex_ssum_defined_iff:
   1.121 + "(\<exists>x. P x \<and> x \<noteq> \<bottom>) =
   1.122 + ((\<exists>x. P (sinl\<cdot>x) \<and> x \<noteq> \<bottom>) \<or>
   1.123 +  (\<exists>x. P (sinr\<cdot>x) \<and> x \<noteq> \<bottom>))"
   1.124 + apply (rule iffI)
   1.125 +  apply (erule exE)
   1.126 +  apply (erule conjE)
   1.127 +  apply (rule_tac p=x in ssumE)
   1.128 +    apply simp
   1.129 +   apply (rule disjI1, fast)
   1.130 +  apply (rule disjI2, fast)
   1.131 + apply (erule disjE)
   1.132 +  apply (force intro: defined_sinl)
   1.133 + apply (force intro: defined_sinr)
   1.134 +done
   1.135 +
   1.136 +lemma exh_start: "p = \<bottom> \<or> (\<exists>x. p = x \<and> x \<noteq> \<bottom>)"
   1.137 +by auto
   1.138 +
   1.139 +lemmas ex_defined_iffs =
   1.140 +   ex_ssum_defined_iff
   1.141 +   ex_sprod_up_defined_iff
   1.142 +   ex_sprod_defined_iff
   1.143 +   ex_up_defined_iff
   1.144 +   ex_one_defined_iff
   1.145 +
   1.146 +text {* Rules for turning exh into casedist *}
   1.147 +
   1.148 +lemma exh_casedist0: "\<lbrakk>R; R \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" (* like make_elim *)
   1.149 +by auto
   1.150 +
   1.151 +lemma exh_casedist1: "((P \<or> Q \<Longrightarrow> R) \<Longrightarrow> S) \<equiv> (\<lbrakk>P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> S)"
   1.152 +by rule auto
   1.153 +
   1.154 +lemma exh_casedist2: "(\<exists>x. P x \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)"
   1.155 +by rule auto
   1.156 +
   1.157 +lemma exh_casedist3: "(P \<and> Q \<Longrightarrow> R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> R)"
   1.158 +by rule auto
   1.159 +
   1.160 +lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3
   1.161 +
   1.162 +
   1.163 +subsection {* Setting up the package *}
   1.164 +
   1.165 +ML_setup {*
   1.166 +val iso_intro       = thm "iso.intro";
   1.167 +val iso_abs_iso     = thm "iso.abs_iso";
   1.168 +val iso_rep_iso     = thm "iso.rep_iso";
   1.169 +val iso_abs_strict  = thm "iso.abs_strict";
   1.170 +val iso_rep_strict  = thm "iso.rep_strict";
   1.171 +val iso_abs_defin'  = thm "iso.abs_defin'";
   1.172 +val iso_rep_defin'  = thm "iso.rep_defin'";
   1.173 +val iso_abs_defined = thm "iso.abs_defined";
   1.174 +val iso_rep_defined = thm "iso.rep_defined";
   1.175 +val iso_iso_swap    = thm "iso.iso_swap";
   1.176 +
   1.177 +val exh_start = thm "exh_start";
   1.178 +val ex_defined_iffs = thms "ex_defined_iffs";
   1.179 +val exh_casedist0 = thm "exh_casedist0";
   1.180 +val exh_casedists = thms "exh_casedists";
   1.181 +*}
   1.182 +
   1.183 +end