src/HOLCF/Representable.thy
author wenzelm
Mon, 25 Oct 2010 21:23:09 +0200
changeset 40133 b61d52de66f0
parent 40038 9d061b3d8f46
child 40216 366309dfaf60
permissions -rw-r--r--
more explicit indication of Output.Private_Hooks (still public to accomodate clones of Proof General for now);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33589
e7ba88cdf3a2 add title/author block
huffman
parents: 33588
diff changeset
     1
(*  Title:      HOLCF/Representable.thy
e7ba88cdf3a2 add title/author block
huffman
parents: 33588
diff changeset
     2
    Author:     Brian Huffman
e7ba88cdf3a2 add title/author block
huffman
parents: 33588
diff changeset
     3
*)
e7ba88cdf3a2 add title/author block
huffman
parents: 33588
diff changeset
     4
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
     5
header {* Representable Types *}
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
     6
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
     7
theory Representable
39987
8c2f449af35a move all bifinite class instances to Bifinite.thy
huffman
parents: 39986
diff changeset
     8
imports Algebraic Bifinite Universal Ssum One Fixrec Domain_Aux
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
     9
uses
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
    10
  ("Tools/repdef.ML")
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
    11
  ("Tools/Domain/domain_isomorphism.ML")
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    12
begin
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    13
39986
38677db30cad rename class 'sfp' to 'bifinite'
huffman
parents: 39974
diff changeset
    14
default_sort bifinite
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    15
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    16
subsection {* Representations of types *}
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    17
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
    18
lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a) = cast\<cdot>DEFL('a)\<cdot>x"
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
    19
by (simp add: cast_DEFL)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    20
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    21
lemma emb_prj_emb:
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    22
  fixes x :: "'a"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    23
  assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    24
  shows "emb\<cdot>(prj\<cdot>(emb\<cdot>x) :: 'b) = emb\<cdot>x"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    25
unfolding emb_prj
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    26
apply (rule cast.belowD)
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    27
apply (rule monofun_cfun_arg [OF assms])
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    28
apply (simp add: cast_DEFL)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    29
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    30
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    31
lemma prj_emb_prj:
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    32
  assumes "DEFL('a) \<sqsubseteq> DEFL('b)"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    33
  shows "prj\<cdot>(emb\<cdot>(prj\<cdot>x :: 'b)) = (prj\<cdot>x :: 'a)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    34
 apply (rule emb_eq_iff [THEN iffD1])
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    35
 apply (simp only: emb_prj)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    36
 apply (rule deflation_below_comp1)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    37
   apply (rule deflation_cast)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    38
  apply (rule deflation_cast)
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    39
 apply (rule monofun_cfun_arg [OF assms])
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    40
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    41
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    42
text {* Isomorphism lemmas used internally by the domain package: *}
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    43
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    44
lemma domain_abs_iso:
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    45
  fixes abs and rep
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
    46
  assumes DEFL: "DEFL('b) = DEFL('a)"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    47
  assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    48
  assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    49
  shows "rep\<cdot>(abs\<cdot>x) = x"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    50
unfolding abs_def rep_def
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    51
by (simp add: emb_prj_emb DEFL)
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    52
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    53
lemma domain_rep_iso:
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    54
  fixes abs and rep
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
    55
  assumes DEFL: "DEFL('b) = DEFL('a)"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    56
  assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    57
  assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    58
  shows "abs\<cdot>(rep\<cdot>x) = x"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    59
unfolding abs_def rep_def
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    60
by (simp add: emb_prj_emb DEFL)
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
    61
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    62
subsection {* Deflations as sets *}
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    63
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    64
definition defl_set :: "defl \<Rightarrow> udom set"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    65
where "defl_set A = {x. cast\<cdot>A\<cdot>x = x}"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    66
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    67
lemma adm_defl_set: "adm (\<lambda>x. x \<in> defl_set A)"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    68
unfolding defl_set_def by simp
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    69
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    70
lemma defl_set_bottom: "\<bottom> \<in> defl_set A"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    71
unfolding defl_set_def by simp
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    72
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    73
lemma defl_set_cast [simp]: "cast\<cdot>A\<cdot>x \<in> defl_set A"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    74
unfolding defl_set_def by simp
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    75
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    76
lemma defl_set_subset_iff: "defl_set A \<subseteq> defl_set B \<longleftrightarrow> A \<sqsubseteq> B"
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    77
apply (simp add: defl_set_def subset_eq cast_below_cast [symmetric])
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    78
apply (auto simp add: cast.belowI cast.belowD)
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    79
done
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    80
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    81
subsection {* Proving a subtype is representable *}
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    82
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    83
text {*
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    84
  Temporarily relax type constraints for @{term emb} and @{term prj}.
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    85
*}
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    86
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    87
setup {*
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    88
  fold Sign.add_const_constraint
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    89
  [ (@{const_name defl}, SOME @{typ "'a::pcpo itself \<Rightarrow> defl"})
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    90
  , (@{const_name emb}, SOME @{typ "'a::pcpo \<rightarrow> udom"})
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    91
  , (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::pcpo"}) ]
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
    92
*}
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    93
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    94
lemma typedef_rep_class:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    95
  fixes Rep :: "'a::pcpo \<Rightarrow> udom"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    96
  fixes Abs :: "udom \<Rightarrow> 'a::pcpo"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
    97
  fixes t :: defl
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
    98
  assumes type: "type_definition Rep Abs (defl_set t)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
    99
  assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
33679
331712879666 automate definition of representable domains from algebraic deflations
huffman
parents: 33589
diff changeset
   100
  assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"
331712879666 automate definition of representable domains from algebraic deflations
huffman
parents: 33589
diff changeset
   101
  assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   102
  assumes defl: "defl \<equiv> (\<lambda> a::'a itself. t)"
39986
38677db30cad rename class 'sfp' to 'bifinite'
huffman
parents: 39974
diff changeset
   103
  shows "OFCLASS('a, bifinite_class)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   104
proof
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   105
  have emb_beta: "\<And>x. emb\<cdot>x = Rep x"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   106
    unfolding emb
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   107
    apply (rule beta_cfun)
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
   108
    apply (rule typedef_cont_Rep [OF type below adm_defl_set])
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   109
    done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   110
  have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   111
    unfolding prj
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   112
    apply (rule beta_cfun)
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
   113
    apply (rule typedef_cont_Abs [OF type below adm_defl_set])
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   114
    apply simp_all
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   115
    done
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
   116
  have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   117
    using type_definition.Rep [OF type]
40038
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
   118
    unfolding prj_beta emb_beta defl_set_def
9d061b3d8f46 replace 'in_defl' relation and '_ ::: _' syntax with 'defl_set' function
huffman
parents: 40037
diff changeset
   119
    by (simp add: type_definition.Rep_inverse [OF type])
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   120
  have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   121
    unfolding prj_beta emb_beta
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   122
    by (simp add: type_definition.Abs_inverse [OF type])
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   123
  show "ep_pair (emb :: 'a \<rightarrow> udom) prj"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   124
    apply default
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   125
    apply (simp add: prj_emb)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   126
    apply (simp add: emb_prj cast.below)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   127
    done
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   128
  show "cast\<cdot>DEFL('a) = emb oo (prj :: udom \<rightarrow> 'a)"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39989
diff changeset
   129
    by (rule cfun_eqI, simp add: defl emb_prj)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   130
qed
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   131
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   132
lemma typedef_DEFL:
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   133
  assumes "defl \<equiv> (\<lambda>a::'a::pcpo itself. t)"
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   134
  shows "DEFL('a::pcpo) = t"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   135
unfolding assms ..
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   136
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   137
text {* Restore original typing constraints. *}
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   138
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   139
setup {*
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   140
  fold Sign.add_const_constraint
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   141
  [ (@{const_name defl}, SOME @{typ "'a::bifinite itself \<Rightarrow> defl"})
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   142
  , (@{const_name emb}, SOME @{typ "'a::bifinite \<rightarrow> udom"})
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   143
  , (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::bifinite"}) ]
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   144
*}
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   145
33679
331712879666 automate definition of representable domains from algebraic deflations
huffman
parents: 33589
diff changeset
   146
use "Tools/repdef.ML"
331712879666 automate definition of representable domains from algebraic deflations
huffman
parents: 33589
diff changeset
   147
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   148
subsection {* Isomorphic deflations *}
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   149
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   150
definition
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   151
  isodefl :: "('a \<rightarrow> 'a) \<Rightarrow> defl \<Rightarrow> bool"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   152
where
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   153
  "isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   154
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   155
lemma isodeflI: "(\<And>x. cast\<cdot>t\<cdot>x = emb\<cdot>(d\<cdot>(prj\<cdot>x))) \<Longrightarrow> isodefl d t"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39989
diff changeset
   156
unfolding isodefl_def by (simp add: cfun_eqI)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   157
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   158
lemma cast_isodefl: "isodefl d t \<Longrightarrow> cast\<cdot>t = (\<Lambda> x. emb\<cdot>(d\<cdot>(prj\<cdot>x)))"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39989
diff changeset
   159
unfolding isodefl_def by (simp add: cfun_eqI)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   160
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   161
lemma isodefl_strict: "isodefl d t \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   162
unfolding isodefl_def
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   163
by (drule cfun_fun_cong [where x="\<bottom>"], simp)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   164
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   165
lemma isodefl_imp_deflation:
39986
38677db30cad rename class 'sfp' to 'bifinite'
huffman
parents: 39974
diff changeset
   166
  fixes d :: "'a \<rightarrow> 'a"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   167
  assumes "isodefl d t" shows "deflation d"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   168
proof
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   169
  note assms [unfolded isodefl_def, simp]
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   170
  fix x :: 'a
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   171
  show "d\<cdot>(d\<cdot>x) = d\<cdot>x"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   172
    using cast.idem [of t "emb\<cdot>x"] by simp
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   173
  show "d\<cdot>x \<sqsubseteq> x"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   174
    using cast.below [of t "emb\<cdot>x"] by simp
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   175
qed
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   176
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   177
lemma isodefl_ID_DEFL: "isodefl (ID :: 'a \<rightarrow> 'a) DEFL('a)"
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   178
unfolding isodefl_def by (simp add: cast_DEFL)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   179
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   180
lemma isodefl_DEFL_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) DEFL('a) \<Longrightarrow> d = ID"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   181
unfolding isodefl_def
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   182
apply (simp add: cast_DEFL)
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39989
diff changeset
   183
apply (simp add: cfun_eq_iff)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   184
apply (rule allI)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   185
apply (drule_tac x="emb\<cdot>x" in spec)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   186
apply simp
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   187
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   188
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   189
lemma isodefl_bottom: "isodefl \<bottom> \<bottom>"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39989
diff changeset
   190
unfolding isodefl_def by (simp add: cfun_eq_iff)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   191
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   192
lemma adm_isodefl:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   193
  "cont f \<Longrightarrow> cont g \<Longrightarrow> adm (\<lambda>x. isodefl (f x) (g x))"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   194
unfolding isodefl_def by simp
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   195
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   196
lemma isodefl_lub:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   197
  assumes "chain d" and "chain t"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   198
  assumes "\<And>i. isodefl (d i) (t i)"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   199
  shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   200
using prems unfolding isodefl_def
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   201
by (simp add: contlub_cfun_arg contlub_cfun_fun)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   202
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   203
lemma isodefl_fix:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   204
  assumes "\<And>d t. isodefl d t \<Longrightarrow> isodefl (f\<cdot>d) (g\<cdot>t)"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   205
  shows "isodefl (fix\<cdot>f) (fix\<cdot>g)"
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   206
unfolding fix_def2
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   207
apply (rule isodefl_lub, simp, simp)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   208
apply (induct_tac i)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   209
apply (simp add: isodefl_bottom)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   210
apply (simp add: assms)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   211
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   212
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
   213
lemma isodefl_abs_rep:
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
   214
  fixes abs and rep and d
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   215
  assumes DEFL: "DEFL('b) = DEFL('a)"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   216
  assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb"
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   217
  assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb"
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
   218
  shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t"
40037
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   219
unfolding isodefl_def
81e6b89d8f58 eliminate constant 'coerce'; use 'prj oo emb' instead
huffman
parents: 40002
diff changeset
   220
by (simp add: cfun_eq_iff assms prj_emb_prj emb_prj_emb)
33779
b8efeea2cebd remove one_typ and tr_typ; add abs/rep lemmas
huffman
parents: 33679
diff changeset
   221
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   222
lemma isodefl_cfun:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   223
  "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   224
    isodefl (cfun_map\<cdot>d1\<cdot>d2) (cfun_defl\<cdot>t1\<cdot>t2)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   225
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   226
apply (simp add: cast_cfun_defl cast_isodefl)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   227
apply (simp add: emb_cfun_def prj_cfun_def)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   228
apply (simp add: cfun_map_map cfcomp1)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   229
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   230
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   231
lemma isodefl_ssum:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   232
  "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   233
    isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   234
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   235
apply (simp add: cast_ssum_defl cast_isodefl)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   236
apply (simp add: emb_ssum_def prj_ssum_def)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   237
apply (simp add: ssum_map_map isodefl_strict)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   238
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   239
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   240
lemma isodefl_sprod:
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   241
  "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   242
    isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   243
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   244
apply (simp add: cast_sprod_defl cast_isodefl)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   245
apply (simp add: emb_sprod_def prj_sprod_def)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   246
apply (simp add: sprod_map_map isodefl_strict)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   247
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   248
33786
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   249
lemma isodefl_cprod:
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   250
  "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   251
    isodefl (cprod_map\<cdot>d1\<cdot>d2) (prod_defl\<cdot>t1\<cdot>t2)"
33786
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   252
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   253
apply (simp add: cast_prod_defl cast_isodefl)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 37770
diff changeset
   254
apply (simp add: emb_prod_def prj_prod_def)
33786
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   255
apply (simp add: cprod_map_map cfcomp1)
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   256
done
d280c5ebd7d7 add lemma isodefl_cprod
huffman
parents: 33784
diff changeset
   257
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   258
lemma isodefl_u:
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   259
  "isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)"
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   260
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   261
apply (simp add: cast_u_defl cast_isodefl)
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   262
apply (simp add: emb_u_def prj_u_def)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   263
apply (simp add: u_map_map)
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   264
done
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   265
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   266
subsection {* Constructing Domain Isomorphisms *}
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   267
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   268
use "Tools/Domain/domain_isomorphism.ML"
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   269
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   270
setup {*
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   271
  fold Domain_Isomorphism.add_type_constructor
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   272
    [(@{type_name cfun}, @{term cfun_defl}, @{const_name cfun_map}, @{thm DEFL_cfun},
35479
dffffe36344a store deflation thms for map functions in theory data
huffman
parents: 35475
diff changeset
   273
        @{thm isodefl_cfun}, @{thm cfun_map_ID}, @{thm deflation_cfun_map}),
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   274
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   275
     (@{type_name ssum}, @{term ssum_defl}, @{const_name ssum_map}, @{thm DEFL_ssum},
35479
dffffe36344a store deflation thms for map functions in theory data
huffman
parents: 35475
diff changeset
   276
        @{thm isodefl_ssum}, @{thm ssum_map_ID}, @{thm deflation_ssum_map}),
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   277
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   278
     (@{type_name sprod}, @{term sprod_defl}, @{const_name sprod_map}, @{thm DEFL_sprod},
35479
dffffe36344a store deflation thms for map functions in theory data
huffman
parents: 35475
diff changeset
   279
        @{thm isodefl_sprod}, @{thm sprod_map_ID}, @{thm deflation_sprod_map}),
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   280
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   281
     (@{type_name prod}, @{term cprod_defl}, @{const_name cprod_map}, @{thm DEFL_prod},
35479
dffffe36344a store deflation thms for map functions in theory data
huffman
parents: 35475
diff changeset
   282
        @{thm isodefl_cprod}, @{thm cprod_map_ID}, @{thm deflation_cprod_map}),
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   283
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39987
diff changeset
   284
     (@{type_name "u"}, @{term u_defl}, @{const_name u_map}, @{thm DEFL_u},
35479
dffffe36344a store deflation thms for map functions in theory data
huffman
parents: 35475
diff changeset
   285
        @{thm isodefl_u}, @{thm u_map_ID}, @{thm deflation_u_map})]
33794
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   286
*}
364bc92ba081 set up domain_isomorphism package in Representable.thy
huffman
parents: 33786
diff changeset
   287
33588
ea9becc59636 theory of representable cpos
huffman
parents:
diff changeset
   288
end