src/HOLCF/Algebraic.thy
author huffman
Wed, 10 Nov 2010 17:56:08 -0800
changeset 40502 8e92772bc0e8
parent 40327 1dfdbd66093a
permissions -rw-r--r--
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/Algebraic.thy
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     3
*)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     4
39985
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39984
diff changeset
     5
header {* Algebraic deflations *}
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     6
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     7
theory Algebraic
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents: 40327
diff changeset
     8
imports Universal Map_Functions
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
     9
begin
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    10
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    11
subsection {* Type constructor for finite deflations *}
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    12
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    13
typedef (open) fin_defl = "{d::udom \<rightarrow> udom. finite_deflation d}"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    14
by (fast intro: finite_deflation_UU)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    15
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    16
instantiation fin_defl :: below
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    17
begin
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    18
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    19
definition below_fin_defl_def:
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    20
    "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep_fin_defl x \<sqsubseteq> Rep_fin_defl y"
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    21
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    22
instance ..
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    23
end
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    24
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    25
instance fin_defl :: po
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    26
using type_definition_fin_defl below_fin_defl_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    27
by (rule typedef_po)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    28
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    29
lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)"
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    30
using Rep_fin_defl by simp
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    31
31164
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
    32
lemma deflation_Rep_fin_defl: "deflation (Rep_fin_defl d)"
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
    33
using finite_deflation_Rep_fin_defl
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
    34
by (rule finite_deflation_imp_deflation)
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
    35
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 29252
diff changeset
    36
interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    37
by (rule finite_deflation_Rep_fin_defl)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    38
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    39
lemma fin_defl_belowI:
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    40
  "(\<And>x. Rep_fin_defl a\<cdot>x = x \<Longrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a \<sqsubseteq> b"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    41
unfolding below_fin_defl_def
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    42
by (rule Rep_fin_defl.belowI)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    43
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    44
lemma fin_defl_belowD:
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    45
  "\<lbrakk>a \<sqsubseteq> b; Rep_fin_defl a\<cdot>x = x\<rbrakk> \<Longrightarrow> Rep_fin_defl b\<cdot>x = x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    46
unfolding below_fin_defl_def
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    47
by (rule Rep_fin_defl.belowD)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    48
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    49
lemma fin_defl_eqI:
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    50
  "(\<And>x. Rep_fin_defl a\<cdot>x = x \<longleftrightarrow> Rep_fin_defl b\<cdot>x = x) \<Longrightarrow> a = b"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    51
apply (rule below_antisym)
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    52
apply (rule fin_defl_belowI, simp)
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    53
apply (rule fin_defl_belowI, simp)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    54
done
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    55
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    56
lemma Rep_fin_defl_mono: "a \<sqsubseteq> b \<Longrightarrow> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    57
unfolding below_fin_defl_def .
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    58
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    59
lemma Abs_fin_defl_mono:
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    60
  "\<lbrakk>finite_deflation a; finite_deflation b; a \<sqsubseteq> b\<rbrakk>
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    61
    \<Longrightarrow> Abs_fin_defl a \<sqsubseteq> Abs_fin_defl b"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    62
unfolding below_fin_defl_def
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    63
by (simp add: Abs_fin_defl_inverse)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    64
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    65
lemma (in finite_deflation) compact_belowI:
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    66
  assumes "\<And>x. compact x \<Longrightarrow> d\<cdot>x = x \<Longrightarrow> f\<cdot>x = x" shows "d \<sqsubseteq> f"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    67
by (rule belowI, rule assms, erule subst, rule compact)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    68
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    69
lemma compact_Rep_fin_defl [simp]: "compact (Rep_fin_defl a)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    70
using finite_deflation_Rep_fin_defl
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    71
by (rule finite_deflation_imp_compact)
33586
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
    72
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    73
subsection {* Defining algebraic deflations by ideal completion *}
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    74
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    75
typedef (open) defl = "{S::fin_defl set. below.ideal S}"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
    76
by (fast intro: below.ideal_principal)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    77
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    78
instantiation defl :: below
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    79
begin
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    80
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    81
definition
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    82
  "x \<sqsubseteq> y \<longleftrightarrow> Rep_defl x \<subseteq> Rep_defl y"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    83
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    84
instance ..
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    85
end
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    86
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    87
instance defl :: po
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    88
using type_definition_defl below_defl_def
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    89
by (rule below.typedef_ideal_po)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    90
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    91
instance defl :: cpo
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    92
using type_definition_defl below_defl_def
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    93
by (rule below.typedef_ideal_cpo)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    94
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    95
definition
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    96
  defl_principal :: "fin_defl \<Rightarrow> defl" where
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
    97
  "defl_principal t = Abs_defl {u. u \<sqsubseteq> t}"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
    98
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
    99
lemma fin_defl_countable: "\<exists>f::fin_defl \<Rightarrow> nat. inj f"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   100
proof
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   101
  have *: "\<And>d. finite (approx_chain.place udom_approx `
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   102
               Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x})"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   103
    apply (rule finite_imageI)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   104
    apply (rule finite_vimageI)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   105
    apply (rule Rep_fin_defl.finite_fixes)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   106
    apply (simp add: inj_on_def Rep_compact_basis_inject)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   107
    done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   108
  have range_eq: "range Rep_compact_basis = {x. compact x}"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   109
    using type_definition_compact_basis by (rule type_definition.Rep_range)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   110
  show "inj (\<lambda>d. set_encode
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   111
    (approx_chain.place udom_approx ` Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x}))"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   112
    apply (rule inj_onI)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   113
    apply (simp only: set_encode_eq *)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   114
    apply (simp only: inj_image_eq_iff approx_chain.inj_place [OF udom_approx])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   115
    apply (drule_tac f="image Rep_compact_basis" in arg_cong)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   116
    apply (simp del: vimage_Collect_eq add: range_eq set_eq_iff)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   117
    apply (rule Rep_fin_defl_inject [THEN iffD1])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   118
    apply (rule below_antisym)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   119
    apply (rule Rep_fin_defl.compact_belowI, rename_tac z)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   120
    apply (drule_tac x=z in spec, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   121
    apply (rule Rep_fin_defl.compact_belowI, rename_tac z)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   122
    apply (drule_tac x=z in spec, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   123
    done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   124
qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   125
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   126
interpretation defl: ideal_completion below defl_principal Rep_defl
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   127
using type_definition_defl below_defl_def
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   128
using defl_principal_def fin_defl_countable
39984
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   129
by (rule below.typedef_ideal_completion)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   130
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   131
text {* Algebraic deflations are pointed *}
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   132
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   133
lemma defl_minimal: "defl_principal (Abs_fin_defl \<bottom>) \<sqsubseteq> x"
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   134
apply (induct x rule: defl.principal_induct, simp)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   135
apply (rule defl.principal_mono)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   136
apply (simp add: below_fin_defl_def)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   137
apply (simp add: Abs_fin_defl_inverse finite_deflation_UU)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   138
done
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   139
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   140
instance defl :: pcpo
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   141
by intro_classes (fast intro: defl_minimal)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   142
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   143
lemma inst_defl_pcpo: "\<bottom> = defl_principal (Abs_fin_defl \<bottom>)"
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   144
by (rule defl_minimal [THEN UU_I, symmetric])
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   145
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   146
subsection {* Applying algebraic deflations *}
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   147
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   148
definition
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   149
  cast :: "defl \<rightarrow> udom \<rightarrow> udom"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   150
where
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   151
  "cast = defl.basis_fun Rep_fin_defl"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   152
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   153
lemma cast_defl_principal:
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   154
  "cast\<cdot>(defl_principal a) = Rep_fin_defl a"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   155
unfolding cast_def
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   156
apply (rule defl.basis_fun_principal)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   157
apply (simp only: below_fin_defl_def)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   158
done
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   159
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   160
lemma deflation_cast: "deflation (cast\<cdot>d)"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   161
apply (induct d rule: defl.principal_induct)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   162
apply (rule adm_subst [OF _ adm_deflation], simp)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   163
apply (simp add: cast_defl_principal)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   164
apply (rule finite_deflation_imp_deflation)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   165
apply (rule finite_deflation_Rep_fin_defl)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   166
done
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   167
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   168
lemma finite_deflation_cast:
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   169
  "compact d \<Longrightarrow> finite_deflation (cast\<cdot>d)"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   170
apply (drule defl.compact_imp_principal, clarify)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   171
apply (simp add: cast_defl_principal)
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   172
apply (rule finite_deflation_Rep_fin_defl)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   173
done
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   174
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 29252
diff changeset
   175
interpretation cast: deflation "cast\<cdot>d"
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   176
by (rule deflation_cast)
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   177
33586
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   178
declare cast.idem [simp]
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   179
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   180
lemma compact_cast [simp]: "compact d \<Longrightarrow> compact (cast\<cdot>d)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   181
apply (rule finite_deflation_imp_compact)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   182
apply (erule finite_deflation_cast)
31164
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   183
done
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   184
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   185
lemma cast_below_cast: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<longleftrightarrow> A \<sqsubseteq> B"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   186
apply (induct A rule: defl.principal_induct, simp)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   187
apply (induct B rule: defl.principal_induct, simp)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   188
apply (simp add: cast_defl_principal below_fin_defl_def)
31164
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   189
done
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   190
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   191
lemma compact_cast_iff: "compact (cast\<cdot>d) \<longleftrightarrow> compact d"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   192
apply (rule iffI)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   193
apply (simp only: compact_def cast_below_cast [symmetric])
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40038
diff changeset
   194
apply (erule adm_subst [OF cont_Rep_cfun2])
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   195
apply (erule compact_cast)
39972
4244ff4f9649 add lemmas finite_deflation_imp_compact, cast_below_cast_iff
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   196
done
4244ff4f9649 add lemmas finite_deflation_imp_compact, cast_below_cast_iff
Brian Huffman <brianh@cs.pdx.edu>
parents: 39971
diff changeset
   197
31164
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   198
lemma cast_below_imp_below: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<Longrightarrow> A \<sqsubseteq> B"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39972
diff changeset
   199
by (simp only: cast_below_cast)
31164
f550c4cf3f3a continuity proofs for approx function on deflations; lemma cast_below_imp_below
huffman
parents: 31076
diff changeset
   200
33586
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   201
lemma cast_eq_imp_eq: "cast\<cdot>A = cast\<cdot>B \<Longrightarrow> A = B"
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   202
by (simp add: below_antisym cast_below_imp_below)
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   203
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   204
lemma cast_strict1 [simp]: "cast\<cdot>\<bottom> = \<bottom>"
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   205
apply (subst inst_defl_pcpo)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39985
diff changeset
   206
apply (subst cast_defl_principal)
33586
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   207
apply (rule Abs_fin_defl_inverse)
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   208
apply (simp add: finite_deflation_UU)
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   209
done
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   210
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   211
lemma cast_strict2 [simp]: "cast\<cdot>A\<cdot>\<bottom> = \<bottom>"
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   212
by (rule cast.below [THEN UU_I])
0e745228d605 add in_deflation relation, more lemmas about cast
huffman
parents: 31164
diff changeset
   213
27409
f65a889f97f9 theory of algebraic deflations
huffman
parents:
diff changeset
   214
end