src/HOLCF/Lift.thy
author wenzelm
Thu, 11 Dec 2008 16:50:18 +0100
changeset 29063 7619f0561cd7
parent 27311 aa28b1d33866
child 29138 661a8db7e647
permissions -rw-r--r--
pcpodef package: state two goals, instead of encoded conjunction;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
     1
(*  Title:      HOLCF/Lift.thy
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
     2
    ID:         $Id$
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
     3
    Author:     Olaf Mueller
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
     4
*)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
     5
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 12026
diff changeset
     6
header {* Lifting types of class type to flat pcpo's *}
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
     7
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
theory Lift
27311
aa28b1d33866 remove unused constant liftpair
huffman
parents: 27310
diff changeset
     9
imports Discrete Up Countable
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
begin
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    11
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 12026
diff changeset
    12
defaultsort type
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    13
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    14
pcpodef 'a lift = "UNIV :: 'a discr u set"
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 27311
diff changeset
    15
by simp_all
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    16
25827
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25723
diff changeset
    17
instance lift :: (finite) finite_po
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25723
diff changeset
    18
by (rule typedef_finite_po [OF type_definition_lift])
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25723
diff changeset
    19
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    20
lemmas inst_lift_pcpo = Abs_lift_strict [symmetric]
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    21
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
    22
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
    23
  Def :: "'a \<Rightarrow> 'a lift" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
    24
  "Def x = Abs_lift (up\<cdot>(Discr x))"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    25
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    26
subsection {* Lift as a datatype *}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    27
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    28
lemma lift_induct: "\<lbrakk>P \<bottom>; \<And>x. P (Def x)\<rbrakk> \<Longrightarrow> P y"
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    29
apply (induct y)
16755
fd02f9d06e43 renamed upE1 to upE; added simp rule cont2cont_flift1
huffman
parents: 16749
diff changeset
    30
apply (rule_tac p=y in upE)
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    31
apply (simp add: Abs_lift_strict)
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    32
apply (case_tac x)
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    33
apply (simp add: Def_def)
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    34
done
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    35
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    36
rep_datatype "\<bottom>\<Colon>'a lift" Def
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    37
  by (erule lift_induct) (simp_all add: Def_def Abs_lift_inject lift_def inst_lift_pcpo)
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    38
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    39
lemmas lift_distinct1 = lift.distinct(1)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    40
lemmas lift_distinct2 = lift.distinct(2)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    41
lemmas Def_not_UU = lift.distinct(2)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 26963
diff changeset
    42
lemmas Def_inject = lift.inject
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    43
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    44
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    45
text {* @{term UU} and @{term Def} *}
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    46
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    47
lemma Lift_exhaust: "x = \<bottom> \<or> (\<exists>y. x = Def y)"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    48
  by (induct x) simp_all
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    49
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    50
lemma Lift_cases: "\<lbrakk>x = \<bottom> \<Longrightarrow> P; \<exists>a. x = Def a \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    51
  by (insert Lift_exhaust) blast
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    52
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    53
lemma not_Undef_is_Def: "(x \<noteq> \<bottom>) = (\<exists>y. x = Def y)"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    54
  by (cases x) simp_all
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    55
16630
83bf468b1dc7 added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents: 16555
diff changeset
    56
lemma lift_definedE: "\<lbrakk>x \<noteq> \<bottom>; \<And>a. x = Def a \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
83bf468b1dc7 added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents: 16555
diff changeset
    57
  by (cases x) simp_all
83bf468b1dc7 added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents: 16555
diff changeset
    58
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    59
text {*
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    60
  For @{term "x ~= UU"} in assumptions @{text def_tac} replaces @{text
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    61
  x} by @{text "Def a"} in conclusion. *}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    62
16121
wenzelm
parents: 16067
diff changeset
    63
ML {*
16630
83bf468b1dc7 added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents: 16555
diff changeset
    64
  local val lift_definedE = thm "lift_definedE"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    65
  in val def_tac = SIMPSET' (fn ss =>
16630
83bf468b1dc7 added theorem lift_definedE; moved cont_if to Cont.thy
huffman
parents: 16555
diff changeset
    66
    etac lift_definedE THEN' asm_simp_tac ss)
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    67
  end;
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    68
*}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    69
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    70
lemma DefE: "Def x = \<bottom> \<Longrightarrow> R"
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    71
  by simp
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    72
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    73
lemma DefE2: "\<lbrakk>x = Def s; x = \<bottom>\<rbrakk> \<Longrightarrow> R"
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    74
  by simp
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    75
27292
huffman
parents: 27104
diff changeset
    76
lemma Def_inject_less_eq: "Def x \<sqsubseteq> Def y \<longleftrightarrow> x = y"
16748
58b9ce4fac54 define lift type using pcpodef; cleaned up
huffman
parents: 16695
diff changeset
    77
by (simp add: less_lift_def Def_def Abs_lift_inverse lift_def)
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    78
27292
huffman
parents: 27104
diff changeset
    79
lemma Def_less_is_eq [simp]: "Def x \<sqsubseteq> y \<longleftrightarrow> Def x = y"
huffman
parents: 27104
diff changeset
    80
by (induct y, simp, simp add: Def_inject_less_eq)
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    81
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    82
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    83
subsection {* Lift is flat *}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    84
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 12026
diff changeset
    85
instance lift :: (type) flat
27292
huffman
parents: 27104
diff changeset
    86
proof
huffman
parents: 27104
diff changeset
    87
  fix x y :: "'a lift"
huffman
parents: 27104
diff changeset
    88
  assume "x \<sqsubseteq> y" thus "x = \<bottom> \<or> x = y"
huffman
parents: 27104
diff changeset
    89
    by (induct x) auto
huffman
parents: 27104
diff changeset
    90
qed
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    91
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    92
text {*
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    93
  \medskip Two specific lemmas for the combination of LCF and HOL
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    94
  terms.
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    95
*}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    96
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
    97
lemma cont_Rep_CFun_app [simp]: "\<lbrakk>cont g; cont f\<rbrakk> \<Longrightarrow> cont(\<lambda>x. ((f x)\<cdot>(g x)) s)"
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 17612
diff changeset
    98
by (rule cont2cont_Rep_CFun [THEN cont2cont_fun])
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
    99
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   100
lemma cont_Rep_CFun_app_app [simp]: "\<lbrakk>cont g; cont f\<rbrakk> \<Longrightarrow> cont(\<lambda>x. ((f x)\<cdot>(g x)) s t)"
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 17612
diff changeset
   101
by (rule cont_Rep_CFun_app [THEN cont2cont_fun])
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
   102
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   103
subsection {* Further operations *}
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   104
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   105
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   106
  flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)"  (binder "FLIFT " 10)  where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   107
  "flift1 = (\<lambda>f. (\<Lambda> x. lift_case \<bottom> f x))"
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   108
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   109
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   110
  flift2 :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a lift \<rightarrow> 'b lift)" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 19764
diff changeset
   111
  "flift2 f = (FLIFT x. Def (f x))"
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   112
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   113
subsection {* Continuity Proofs for flift1, flift2 *}
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   114
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   115
text {* Need the instance of @{text flat}. *}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   116
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   117
lemma cont_lift_case1: "cont (\<lambda>f. lift_case a f x)"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   118
apply (induct x)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   119
apply simp
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   120
apply simp
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 17612
diff changeset
   121
apply (rule cont_id [THEN cont2cont_fun])
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   122
done
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   123
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   124
lemma cont_lift_case2: "cont (\<lambda>x. lift_case \<bottom> f x)"
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   125
apply (rule flatdom_strict2cont)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   126
apply simp
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   127
done
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   128
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   129
lemma cont_flift1: "cont flift1"
27292
huffman
parents: 27104
diff changeset
   130
unfolding flift1_def
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   131
apply (rule cont2cont_LAM)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   132
apply (rule cont_lift_case2)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   133
apply (rule cont_lift_case1)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   134
done
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   135
27310
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   136
lemma FLIFT_mono:
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   137
  "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> (FLIFT x. f x) \<sqsubseteq> (FLIFT x. g x)"
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   138
apply (rule monofunE [where f=flift1])
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   139
apply (rule cont2mono [OF cont_flift1])
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   140
apply (simp add: less_fun_ext)
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   141
done
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   142
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   143
lemma cont2cont_flift1 [simp]:
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   144
  "\<lbrakk>\<And>y. cont (\<lambda>x. f x y)\<rbrakk> \<Longrightarrow> cont (\<lambda>x. FLIFT y. f x y)"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   145
apply (rule cont_flift1 [THEN cont2cont_app3])
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   146
apply simp
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   147
done
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   148
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   149
lemma cont2cont_lift_case [simp]:
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   150
  "\<lbrakk>\<And>y. cont (\<lambda>x. f x y); cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. lift_case UU (f x) (g x))"
16757
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   151
apply (subgoal_tac "cont (\<lambda>x. (FLIFT y. f x y)\<cdot>(g x))")
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   152
apply (simp add: flift1_def cont_lift_case2)
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   153
apply simp
16757
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   154
done
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   155
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   156
text {* rewrites for @{term flift1}, @{term flift2} *}
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   157
16695
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   158
lemma flift1_Def [simp]: "flift1 f\<cdot>(Def x) = (f x)"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   159
by (simp add: flift1_def cont_lift_case2)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   160
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   161
lemma flift2_Def [simp]: "flift2 f\<cdot>(Def x) = Def (f x)"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   162
by (simp add: flift2_def)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   163
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   164
lemma flift1_strict [simp]: "flift1 f\<cdot>\<bottom> = \<bottom>"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   165
by (simp add: flift1_def cont_lift_case2)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   166
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   167
lemma flift2_strict [simp]: "flift2 f\<cdot>\<bottom> = \<bottom>"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   168
by (simp add: flift2_def)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   169
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   170
lemma flift2_defined [simp]: "x \<noteq> \<bottom> \<Longrightarrow> (flift2 f)\<cdot>x \<noteq> \<bottom>"
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   171
by (erule lift_definedE, simp)
dc8c868e910b simplified definitions of flift1, flift2, liftpair;
huffman
parents: 16630
diff changeset
   172
19520
873dad2d8a6d add theorem flift2_defined_iff
huffman
parents: 19440
diff changeset
   173
lemma flift2_defined_iff [simp]: "(flift2 f\<cdot>x = \<bottom>) = (x = \<bottom>)"
873dad2d8a6d add theorem flift2_defined_iff
huffman
parents: 19440
diff changeset
   174
by (cases x, simp_all)
873dad2d8a6d add theorem flift2_defined_iff
huffman
parents: 19440
diff changeset
   175
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   176
text {*
14096
f79d139c7e46 corrected markup text
oheimb
parents: 12338
diff changeset
   177
  \medskip Extension of @{text cont_tac} and installation of simplifier.
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   178
*}
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   179
26452
ed657432b8b9 declare cont_lemmas_ext as simp rules individually
huffman
parents: 25920
diff changeset
   180
lemmas cont_lemmas_ext =
16757
b8bfd086f7d4 replaced old continuity rules with new lemma cont2cont_lift_case
huffman
parents: 16755
diff changeset
   181
  cont2cont_flift1 cont2cont_lift_case cont2cont_lambda
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   182
  cont_Rep_CFun_app cont_Rep_CFun_app_app cont_if
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   183
16121
wenzelm
parents: 16067
diff changeset
   184
ML {*
19764
372065f34795 removed obsolete ML files;
wenzelm
parents: 19520
diff changeset
   185
local
372065f34795 removed obsolete ML files;
wenzelm
parents: 19520
diff changeset
   186
  val cont_lemmas2 = thms "cont_lemmas1" @ thms "cont_lemmas_ext";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19520
diff changeset
   187
  val flift1_def = thm "flift1_def";
372065f34795 removed obsolete ML files;
wenzelm
parents: 19520
diff changeset
   188
in
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   189
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   190
fun cont_tac  i = resolve_tac cont_lemmas2 i;
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   191
fun cont_tacR i = REPEAT (cont_tac i);
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   192
19764
372065f34795 removed obsolete ML files;
wenzelm
parents: 19520
diff changeset
   193
fun cont_tacRs ss i =
17612
5b37787d2d9e adm_tac/cont_tacRs: proper simpset;
wenzelm
parents: 16757
diff changeset
   194
  simp_tac ss i THEN
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   195
  REPEAT (cont_tac i)
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   196
end;
15651
4b393520846e Replaced continuity solver with new continuity simproc. Also removed cont lemmas from simp set, so that the simproc actually gets used.
huffman
parents: 15577
diff changeset
   197
*}
12026
0b1d80ada4ab rep_datatype lift;
wenzelm
parents: 2640
diff changeset
   198
26963
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   199
subsection {* Lifted countable types are bifinite *}
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   200
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   201
instantiation lift :: (countable) bifinite
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   202
begin
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   203
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   204
definition
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   205
  approx_lift_def:
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   206
    "approx = (\<lambda>n. FLIFT x. if to_nat x < n then Def x else \<bottom>)"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   207
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   208
instance proof
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   209
  fix x :: "'a lift"
27310
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   210
  show "chain (approx :: nat \<Rightarrow> 'a lift \<rightarrow> 'a lift)"
26963
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   211
    unfolding approx_lift_def
27310
d0229bc6c461 simplify profinite class axioms
huffman
parents: 27292
diff changeset
   212
    by (rule chainI, simp add: FLIFT_mono)
26963
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   213
next
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   214
  fix x :: "'a lift"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   215
  show "(\<Squnion>i. approx i\<cdot>x) = x"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   216
    unfolding approx_lift_def
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   217
    apply (cases x, simp)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   218
    apply (rule thelubI)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   219
    apply (rule is_lubI)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   220
     apply (rule ub_rangeI, simp)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   221
    apply (drule ub_rangeD)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   222
    apply (erule rev_trans_less)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   223
    apply simp
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   224
    apply (rule lessI)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   225
    done
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   226
next
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   227
  fix i :: nat and x :: "'a lift"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   228
  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   229
    unfolding approx_lift_def
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   230
    by (cases x, simp, simp)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   231
next
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   232
  fix i :: nat
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   233
  show "finite {x::'a lift. approx i\<cdot>x = x}"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   234
  proof (rule finite_subset)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   235
    let ?S = "insert (\<bottom>::'a lift) (Def ` to_nat -` {..<i})"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   236
    show "{x::'a lift. approx i\<cdot>x = x} \<subseteq> ?S"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   237
      unfolding approx_lift_def
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   238
      by (rule subsetI, case_tac x, simp, simp split: split_if_asm)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   239
    show "finite ?S"
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   240
      by (simp add: finite_vimageI)
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   241
  qed
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   242
qed
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   243
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents:
diff changeset
   244
end
26963
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   245
290d23780204 instantiation lift :: (countable) bifinite
huffman
parents: 26452
diff changeset
   246
end