src/HOLCF/TypedefPcpo.thy
author huffman
Tue, 24 May 2005 05:51:06 +0200
changeset 16058 3d50b521ab16
child 16070 4a83dd540b88
permissions -rw-r--r--
New theory for defining subtypes of pcpos
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16058
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/TypedefPcpo.thy
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     2
    ID:         $Id$
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     3
    Author:     Brian Huffman
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     5
*)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     6
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     7
header {* Subtypes of pcpos *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     8
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
     9
theory TypedefPcpo
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    10
imports Adm
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    11
begin
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    12
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    13
subsection {* Proving a subtype is a partial order *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    14
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    15
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    16
  A subtype of a partial order is itself a partial order,
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    17
  if the ordering is defined in the standard way.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    18
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    19
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    20
theorem typedef_po:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    21
fixes Abs :: "'a::po \<Rightarrow> 'b::sq_ord"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    22
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    23
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    24
shows "OFCLASS('b, po_class)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    25
 apply (intro_classes, unfold less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    26
   apply (rule refl_less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    27
  apply (subst type_definition.Rep_inject [OF type, symmetric])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    28
  apply (rule antisym_less, assumption+)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    29
 apply (rule trans_less, assumption+)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    30
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    31
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    32
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    33
subsection {* Proving a subtype is complete *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    34
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    35
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    36
  A subtype of a cpo is itself a cpo if the ordering is
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    37
  defined in the standard way, and the defining subset
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    38
  is closed with respect to limits of chains.  A set is
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    39
  closed if and only if membership in the set is an
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    40
  admissible predicate.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    41
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    42
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    43
lemma chain_Rep:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    44
assumes less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    45
shows "chain S \<Longrightarrow> chain (\<lambda>n. Rep (S n))"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    46
by (rule chainI, drule chainE, unfold less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    47
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    48
lemma lub_Rep_in_A:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    49
fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    50
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    51
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    52
    and adm:  "adm (\<lambda>x. x \<in> A)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    53
shows "chain S \<Longrightarrow> (LUB n. Rep (S n)) \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    54
 apply (erule admD [OF adm chain_Rep [OF less], rule_format])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    55
 apply (rule type_definition.Rep [OF type])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    56
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    57
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    58
theorem typedef_is_lub:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    59
fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    60
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    61
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    62
    and adm: "adm (\<lambda>x. x \<in> A)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    63
shows "chain S \<Longrightarrow> range S <<| Abs (LUB n. Rep (S n))"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    64
 apply (rule is_lubI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    65
  apply (rule ub_rangeI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    66
  apply (subst less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    67
  apply (subst type_definition.Abs_inverse [OF type])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    68
   apply (erule lub_Rep_in_A [OF type less adm])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    69
  apply (rule is_ub_thelub)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    70
  apply (erule chain_Rep [OF less])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    71
 apply (subst less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    72
 apply (subst type_definition.Abs_inverse [OF type])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    73
  apply (erule lub_Rep_in_A [OF type less adm])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    74
 apply (rule is_lub_thelub)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    75
  apply (erule chain_Rep [OF less])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    76
 apply (rule ub_rangeI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    77
 apply (drule ub_rangeD)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    78
 apply (unfold less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    79
 apply assumption
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    80
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    81
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    82
theorem typedef_cpo:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    83
fixes Abs :: "'a::cpo \<Rightarrow> 'b::po"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    84
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    85
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    86
    and adm: "adm (\<lambda>x. x \<in> A)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    87
shows "OFCLASS('b, cpo_class)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    88
 apply (intro_classes)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    89
 apply (rule_tac x="Abs (LUB n. Rep (S n))" in exI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    90
 apply (erule typedef_is_lub [OF type less adm])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    91
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    92
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    93
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    94
subsubsection {* Continuity of @{term Rep} and @{term Abs} *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    95
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    96
text {* For any sub-cpo, the @{term Rep} function is continuous. *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    97
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    98
theorem typedef_cont_Rep:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
    99
fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   100
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   101
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   102
    and adm: "adm (\<lambda>x. x \<in> A)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   103
shows "cont Rep"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   104
 apply (rule contI[rule_format])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   105
 apply (simp only: typedef_is_lub [OF type less adm, THEN thelubI])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   106
 apply (subst type_definition.Abs_inverse [OF type])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   107
  apply (erule lub_Rep_in_A [OF type less adm])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   108
 apply (rule thelubE [OF _ refl])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   109
 apply (erule chain_Rep [OF less])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   110
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   111
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   112
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   113
  For a sub-cpo, we can make the @{term Abs} function continuous
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   114
  only if we restrict its domain to the defining subset by
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   115
  composing it with another continuous function.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   116
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   117
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   118
theorem typedef_cont_Abs:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   119
fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   120
fixes f :: "'c::cpo \<Rightarrow> 'a::cpo"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   121
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   122
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   123
    and adm: "adm (\<lambda>x. x \<in> A)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   124
    and f_in_A: "\<And>x. f x \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   125
    and cont_f: "cont f"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   126
shows "cont (\<lambda>x. Abs (f x))"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   127
 apply (rule contI[rule_format])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   128
 apply (rule is_lubI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   129
  apply (rule ub_rangeI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   130
  apply (simp only: less type_definition.Abs_inverse [OF type f_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   131
  apply (rule monofun_fun_arg [OF cont2mono [OF cont_f]])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   132
  apply (erule is_ub_thelub)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   133
 apply (simp only: less type_definition.Abs_inverse [OF type f_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   134
 apply (simp only: contlubE[rule_format, OF cont2contlub [OF cont_f]])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   135
 apply (rule is_lub_thelub)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   136
  apply (erule ch2ch_monofun [OF cont2mono [OF cont_f]])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   137
 apply (rule ub_rangeI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   138
 apply (drule_tac i=i in ub_rangeD)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   139
 apply (simp only: less type_definition.Abs_inverse [OF type f_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   140
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   141
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   142
lemmas typedef_cont_Abs2 =
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   143
  typedef_cont_Abs [OF _ _ _ _ cont_Rep_CFun2]
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   144
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   145
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   146
subsection {* Proving a typedef is pointed *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   147
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   148
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   149
  A subtype of a cpo has a least element if and only if
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   150
  the defining subset has a least element.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   151
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   152
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   153
theorem typedef_pcpo:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   154
fixes Abs :: "'a::cpo \<Rightarrow> 'b::cpo"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   155
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   156
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   157
    and z_in_A: "z \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   158
    and z_least: "\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   159
shows "OFCLASS('b, pcpo_class)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   160
 apply (intro_classes)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   161
 apply (rule_tac x="Abs z" in exI, rule allI)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   162
 apply (unfold less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   163
 apply (subst type_definition.Abs_inverse [OF type z_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   164
 apply (rule z_least [OF type_definition.Rep [OF type]])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   165
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   166
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   167
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   168
  As a special case, a subtype of a pcpo has a least element
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   169
  if the defining subset contains @{term \<bottom>}.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   170
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   171
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   172
theorem typedef_pcpo_UU:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   173
fixes Abs :: "'a::pcpo \<Rightarrow> 'b::cpo"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   174
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   175
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   176
    and UU_in_A: "\<bottom> \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   177
shows "OFCLASS('b, pcpo_class)"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   178
by (rule typedef_pcpo [OF type less UU_in_A], rule minimal)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   179
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   180
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   181
subsubsection {* Strictness of @{term Rep} and @{term Abs} *}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   182
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   183
text {*
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   184
  For a sub-pcpo where @{term \<bottom>} is a member of the defining
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   185
  subset, @{term Rep} and @{term Abs} are both strict.
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   186
*}
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   187
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   188
theorem typedef_strict_Abs:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   189
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   190
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   191
    and UU_in_A: "\<bottom> \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   192
shows "Abs \<bottom> = \<bottom>"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   193
 apply (rule UU_I, unfold less)
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   194
 apply (simp add: type_definition.Abs_inverse [OF type UU_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   195
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   196
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   197
theorem typedef_strict_Rep:
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   198
assumes type: "type_definition Rep Abs A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   199
    and less: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   200
    and UU_in_A: "\<bottom> \<in> A"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   201
shows "Rep \<bottom> = \<bottom>"
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   202
 apply (rule typedef_strict_Abs [OF type less UU_in_A, THEN subst])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   203
 apply (rule type_definition.Abs_inverse [OF type UU_in_A])
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   204
done
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   205
3d50b521ab16 New theory for defining subtypes of pcpos
huffman
parents:
diff changeset
   206
end