src/HOLCF/Domain.thy
author huffman
Fri, 10 Apr 2009 11:35:21 -0700
changeset 30910 a7cc0ef93269
parent 29138 661a8db7e647
child 30911 7809cbaa1b61
permissions -rw-r--r--
set up domain package in Domain.thy
     1 (*  Title:      HOLCF/Domain.thy
     2     Author:     Brian Huffman
     3 *)
     4 
     5 header {* Domain package *}
     6 
     7 theory Domain
     8 imports Ssum Sprod Up One Tr Fixrec
     9 uses
    10   ("Tools/cont_consts.ML")
    11   ("Tools/cont_proc.ML")
    12   ("Tools/domain/domain_library.ML")
    13   ("Tools/domain/domain_syntax.ML")
    14   ("Tools/domain/domain_axioms.ML")
    15   ("Tools/domain/domain_theorems.ML")
    16   ("Tools/domain/domain_extender.ML")
    17 begin
    18 
    19 defaultsort pcpo
    20 
    21 
    22 subsection {* Continuous isomorphisms *}
    23 
    24 text {* A locale for continuous isomorphisms *}
    25 
    26 locale iso =
    27   fixes abs :: "'a \<rightarrow> 'b"
    28   fixes rep :: "'b \<rightarrow> 'a"
    29   assumes abs_iso [simp]: "rep\<cdot>(abs\<cdot>x) = x"
    30   assumes rep_iso [simp]: "abs\<cdot>(rep\<cdot>y) = y"
    31 begin
    32 
    33 lemma swap: "iso rep abs"
    34   by (rule iso.intro [OF rep_iso abs_iso])
    35 
    36 lemma abs_less: "(abs\<cdot>x \<sqsubseteq> abs\<cdot>y) = (x \<sqsubseteq> y)"
    37 proof
    38   assume "abs\<cdot>x \<sqsubseteq> abs\<cdot>y"
    39   then have "rep\<cdot>(abs\<cdot>x) \<sqsubseteq> rep\<cdot>(abs\<cdot>y)" by (rule monofun_cfun_arg)
    40   then show "x \<sqsubseteq> y" by simp
    41 next
    42   assume "x \<sqsubseteq> y"
    43   then show "abs\<cdot>x \<sqsubseteq> abs\<cdot>y" by (rule monofun_cfun_arg)
    44 qed
    45 
    46 lemma rep_less: "(rep\<cdot>x \<sqsubseteq> rep\<cdot>y) = (x \<sqsubseteq> y)"
    47   by (rule iso.abs_less [OF swap])
    48 
    49 lemma abs_eq: "(abs\<cdot>x = abs\<cdot>y) = (x = y)"
    50   by (simp add: po_eq_conv abs_less)
    51 
    52 lemma rep_eq: "(rep\<cdot>x = rep\<cdot>y) = (x = y)"
    53   by (rule iso.abs_eq [OF swap])
    54 
    55 lemma abs_strict: "abs\<cdot>\<bottom> = \<bottom>"
    56 proof -
    57   have "\<bottom> \<sqsubseteq> rep\<cdot>\<bottom>" ..
    58   then have "abs\<cdot>\<bottom> \<sqsubseteq> abs\<cdot>(rep\<cdot>\<bottom>)" by (rule monofun_cfun_arg)
    59   then have "abs\<cdot>\<bottom> \<sqsubseteq> \<bottom>" by simp
    60   then show ?thesis by (rule UU_I)
    61 qed
    62 
    63 lemma rep_strict: "rep\<cdot>\<bottom> = \<bottom>"
    64   by (rule iso.abs_strict [OF swap])
    65 
    66 lemma abs_defin': "abs\<cdot>x = \<bottom> \<Longrightarrow> x = \<bottom>"
    67 proof -
    68   have "x = rep\<cdot>(abs\<cdot>x)" by simp
    69   also assume "abs\<cdot>x = \<bottom>"
    70   also note rep_strict
    71   finally show "x = \<bottom>" .
    72 qed
    73 
    74 lemma rep_defin': "rep\<cdot>z = \<bottom> \<Longrightarrow> z = \<bottom>"
    75   by (rule iso.abs_defin' [OF swap])
    76 
    77 lemma abs_defined: "z \<noteq> \<bottom> \<Longrightarrow> abs\<cdot>z \<noteq> \<bottom>"
    78   by (erule contrapos_nn, erule abs_defin')
    79 
    80 lemma rep_defined: "z \<noteq> \<bottom> \<Longrightarrow> rep\<cdot>z \<noteq> \<bottom>"
    81   by (rule iso.abs_defined [OF iso.swap]) (rule iso_axioms)
    82 
    83 lemma abs_defined_iff: "(abs\<cdot>x = \<bottom>) = (x = \<bottom>)"
    84   by (auto elim: abs_defin' intro: abs_strict)
    85 
    86 lemma rep_defined_iff: "(rep\<cdot>x = \<bottom>) = (x = \<bottom>)"
    87   by (rule iso.abs_defined_iff [OF iso.swap]) (rule iso_axioms)
    88 
    89 lemma (in iso) compact_abs_rev: "compact (abs\<cdot>x) \<Longrightarrow> compact x"
    90 proof (unfold compact_def)
    91   assume "adm (\<lambda>y. \<not> abs\<cdot>x \<sqsubseteq> y)"
    92   with cont_Rep_CFun2
    93   have "adm (\<lambda>y. \<not> abs\<cdot>x \<sqsubseteq> abs\<cdot>y)" by (rule adm_subst)
    94   then show "adm (\<lambda>y. \<not> x \<sqsubseteq> y)" using abs_less by simp
    95 qed
    96 
    97 lemma compact_rep_rev: "compact (rep\<cdot>x) \<Longrightarrow> compact x"
    98   by (rule iso.compact_abs_rev [OF iso.swap]) (rule iso_axioms)
    99 
   100 lemma compact_abs: "compact x \<Longrightarrow> compact (abs\<cdot>x)"
   101   by (rule compact_rep_rev) simp
   102 
   103 lemma compact_rep: "compact x \<Longrightarrow> compact (rep\<cdot>x)"
   104   by (rule iso.compact_abs [OF iso.swap]) (rule iso_axioms)
   105 
   106 lemma iso_swap: "(x = abs\<cdot>y) = (rep\<cdot>x = y)"
   107 proof
   108   assume "x = abs\<cdot>y"
   109   then have "rep\<cdot>x = rep\<cdot>(abs\<cdot>y)" by simp
   110   then show "rep\<cdot>x = y" by simp
   111 next
   112   assume "rep\<cdot>x = y"
   113   then have "abs\<cdot>(rep\<cdot>x) = abs\<cdot>y" by simp
   114   then show "x = abs\<cdot>y" by simp
   115 qed
   116 
   117 end
   118 
   119 
   120 subsection {* Casedist *}
   121 
   122 lemma ex_one_defined_iff:
   123   "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = P ONE"
   124  apply safe
   125   apply (rule_tac p=x in oneE)
   126    apply simp
   127   apply simp
   128  apply force
   129  done
   130 
   131 lemma ex_up_defined_iff:
   132   "(\<exists>x. P x \<and> x \<noteq> \<bottom>) = (\<exists>x. P (up\<cdot>x))"
   133  apply safe
   134   apply (rule_tac p=x in upE)
   135    apply simp
   136   apply fast
   137  apply (force intro!: up_defined)
   138  done
   139 
   140 lemma ex_sprod_defined_iff:
   141  "(\<exists>y. P y \<and> y \<noteq> \<bottom>) =
   142   (\<exists>x y. (P (:x, y:) \<and> x \<noteq> \<bottom>) \<and> y \<noteq> \<bottom>)"
   143  apply safe
   144   apply (rule_tac p=y in sprodE)
   145    apply simp
   146   apply fast
   147  apply (force intro!: spair_defined)
   148  done
   149 
   150 lemma ex_sprod_up_defined_iff:
   151  "(\<exists>y. P y \<and> y \<noteq> \<bottom>) =
   152   (\<exists>x y. P (:up\<cdot>x, y:) \<and> y \<noteq> \<bottom>)"
   153  apply safe
   154   apply (rule_tac p=y in sprodE)
   155    apply simp
   156   apply (rule_tac p=x in upE)
   157    apply simp
   158   apply fast
   159  apply (force intro!: spair_defined)
   160  done
   161 
   162 lemma ex_ssum_defined_iff:
   163  "(\<exists>x. P x \<and> x \<noteq> \<bottom>) =
   164  ((\<exists>x. P (sinl\<cdot>x) \<and> x \<noteq> \<bottom>) \<or>
   165   (\<exists>x. P (sinr\<cdot>x) \<and> x \<noteq> \<bottom>))"
   166  apply (rule iffI)
   167   apply (erule exE)
   168   apply (erule conjE)
   169   apply (rule_tac p=x in ssumE)
   170     apply simp
   171    apply (rule disjI1, fast)
   172   apply (rule disjI2, fast)
   173  apply (erule disjE)
   174   apply force
   175  apply force
   176  done
   177 
   178 lemma exh_start: "p = \<bottom> \<or> (\<exists>x. p = x \<and> x \<noteq> \<bottom>)"
   179   by auto
   180 
   181 lemmas ex_defined_iffs =
   182    ex_ssum_defined_iff
   183    ex_sprod_up_defined_iff
   184    ex_sprod_defined_iff
   185    ex_up_defined_iff
   186    ex_one_defined_iff
   187 
   188 text {* Rules for turning exh into casedist *}
   189 
   190 lemma exh_casedist0: "\<lbrakk>R; R \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" (* like make_elim *)
   191   by auto
   192 
   193 lemma exh_casedist1: "((P \<or> Q \<Longrightarrow> R) \<Longrightarrow> S) \<equiv> (\<lbrakk>P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> S)"
   194   by rule auto
   195 
   196 lemma exh_casedist2: "(\<exists>x. P x \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)"
   197   by rule auto
   198 
   199 lemma exh_casedist3: "(P \<and> Q \<Longrightarrow> R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> R)"
   200   by rule auto
   201 
   202 lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3
   203 
   204 
   205 subsection {* Installing the domain package *}
   206 
   207 use "Tools/cont_consts.ML"
   208 use "Tools/cont_proc.ML"
   209 use "Tools/domain/domain_library.ML"
   210 use "Tools/domain/domain_syntax.ML"
   211 use "Tools/domain/domain_axioms.ML"
   212 use "Tools/domain/domain_theorems.ML"
   213 use "Tools/domain/domain_extender.ML"
   214 
   215 end