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