src/HOLCF/Cfun.thy
author huffman
Wed, 08 Jun 2005 00:04:38 +0200
changeset 16314 7102a1aaecfd
parent 16209 36ee7f6af79f
child 16386 c6f5ade29608
permissions -rw-r--r--
added theorem injection_less
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15600
a59f07556a8d fixed filename in header
huffman
parents: 15589
diff changeset
     1
(*  Title:      HOLCF/Cfun.thy
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     2
    ID:         $Id$
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     3
    Author:     Franz Regensburger
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     4
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     5
Definition of the type ->  of continuous functions.
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     6
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     7
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     8
header {* The type of continuous functions *}
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     9
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
theory Cfun
16094
a92ee2833938 Use TypedefPcpo theorem for po instance
huffman
parents: 16093
diff changeset
    11
imports TypedefPcpo
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    12
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    13
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    14
defaultsort cpo
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    15
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    16
subsection {* Definition of continuous function type *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    17
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    18
typedef (CFun)  ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    19
by (rule exI, fast intro: cont_const)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    20
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    21
syntax
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    22
  Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    23
                                                (* application *)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    24
  Abs_CFun :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    25
                                                (* abstraction *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    26
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    27
syntax (xsymbols)
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    28
  "->"     :: "[type, type] => type"      ("(_ \<rightarrow>/ _)" [1,0]0)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    29
  "LAM "   :: "[idts, 'a => 'b] => ('a -> 'b)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    30
					("(3\<Lambda>_./ _)" [0, 10] 10)
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    31
  Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\<cdot>_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    32
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    33
syntax (HTML output)
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    34
  Rep_CFun :: "('a -> 'b) => ('a => 'b)" ("(_\<cdot>_)" [999,1000] 999)
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    35
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    36
subsection {* Class instances *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    37
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    38
instance "->"  :: (cpo, cpo) sq_ord ..
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    39
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    40
defs (overloaded)
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    41
  less_cfun_def: "(op \<sqsubseteq>) \<equiv> (\<lambda>f g. Rep_CFun f \<sqsubseteq> Rep_CFun g)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    42
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    43
lemma adm_CFun: "adm (\<lambda>f. f \<in> CFun)"
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    44
by (simp add: CFun_def, rule admI, rule cont_lub_fun)
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    45
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    46
lemma UU_CFun: "\<bottom> \<in> CFun"
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    47
by (simp add: CFun_def inst_fun_pcpo cont_const)
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    48
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    49
instance "->" :: (cpo, cpo) po
16094
a92ee2833938 Use TypedefPcpo theorem for po instance
huffman
parents: 16093
diff changeset
    50
by (rule typedef_po [OF type_definition_CFun less_cfun_def])
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    51
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    52
instance "->" :: (cpo, cpo) cpo
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    53
by (rule typedef_cpo [OF type_definition_CFun less_cfun_def adm_CFun])
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    54
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    55
instance "->" :: (cpo, pcpo) pcpo
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    56
by (rule typedef_pcpo_UU [OF type_definition_CFun less_cfun_def UU_CFun])
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    57
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    58
lemmas cont_Rep_CFun =
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    59
  typedef_cont_Rep [OF type_definition_CFun less_cfun_def adm_CFun]
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    60
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    61
lemmas cont_Abs_CFun = 
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    62
  typedef_cont_Abs [OF type_definition_CFun less_cfun_def adm_CFun]
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    63
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    64
lemmas Rep_CFun_strict =
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    65
  typedef_Rep_strict [OF type_definition_CFun less_cfun_def UU_CFun]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    66
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    67
lemmas Abs_CFun_strict =
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    68
  typedef_Abs_strict [OF type_definition_CFun less_cfun_def UU_CFun]
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    69
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    70
text {* Additional lemma about the isomorphism between
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    71
        @{typ "'a -> 'b"} and @{term CFun} *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    72
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    73
lemma Abs_CFun_inverse2: "cont f \<Longrightarrow> Rep_CFun (Abs_CFun f) = f"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    74
by (simp add: Abs_CFun_inverse CFun_def)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    75
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    76
text {* Beta-equality for continuous functions *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    77
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    78
lemma beta_cfun [simp]: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    79
by (simp add: Abs_CFun_inverse2)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    80
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    81
text {* Eta-equality for continuous functions *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    82
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    83
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    84
by (rule Rep_CFun_inverse)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    85
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    86
text {* Extensionality for continuous functions *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    87
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    88
lemma ext_cfun: "(\<And>x. f\<cdot>x = g\<cdot>x) \<Longrightarrow> f = g"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    89
by (simp add: Rep_CFun_inject [symmetric] ext)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    90
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    91
text {* lemmas about application of continuous functions *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    92
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    93
lemma cfun_cong: "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> f\<cdot>x = g\<cdot>y"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    94
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    95
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    96
lemma cfun_fun_cong: "f = g \<Longrightarrow> f\<cdot>x = g\<cdot>x"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    97
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    98
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
    99
lemma cfun_arg_cong: "x = y \<Longrightarrow> f\<cdot>x = f\<cdot>y"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   100
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   101
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   102
subsection {* Continuity of application *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   103
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   104
lemma cont_Rep_CFun1: "cont (\<lambda>f. f\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   105
by (rule cont_Rep_CFun [THEN cont2cont_CF1L])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   106
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   107
lemma cont_Rep_CFun2: "cont (\<lambda>x. f\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   108
apply (rule_tac P = "cont" in CollectD)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   109
apply (fold CFun_def)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   110
apply (rule Rep_CFun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   111
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   112
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   113
lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   114
lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   115
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   116
lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   117
lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   118
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   119
lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   120
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   121
text {* contlub, cont properties of @{term Rep_CFun} in each argument *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   122
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   123
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(lub (range Y)) = (\<Squnion>i. f\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   124
by (rule contlub_Rep_CFun2 [THEN contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   125
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   126
lemma cont_cfun_arg: "chain Y \<Longrightarrow> range (\<lambda>i. f\<cdot>(Y i)) <<| f\<cdot>(lub (range Y))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   127
by (rule cont_Rep_CFun2 [THEN contE])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   128
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   129
lemma contlub_cfun_fun: "chain F \<Longrightarrow> lub (range F)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   130
by (rule contlub_Rep_CFun1 [THEN contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   131
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   132
lemma cont_cfun_fun: "chain F \<Longrightarrow> range (\<lambda>i. F i\<cdot>x) <<| lub (range F)\<cdot>x"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   133
by (rule cont_Rep_CFun1 [THEN contE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   134
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   135
text {* Extensionality wrt. @{term "op <<"} in @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   136
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   137
lemma less_cfun_ext: "(\<And>x. f\<cdot>x \<sqsubseteq> g\<cdot>x) \<Longrightarrow> f \<sqsubseteq> g"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   138
by (simp add: less_cfun_def less_fun_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   139
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   140
text {* monotonicity of application *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   141
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   142
lemma monofun_cfun_fun: "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>x"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   143
by (simp add: less_cfun_def less_fun_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   144
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   145
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   146
by (rule monofun_Rep_CFun2 [THEN monofunE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   147
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   148
lemma monofun_cfun: "\<lbrakk>f \<sqsubseteq> g; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>y"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   149
by (rule trans_less [OF monofun_cfun_fun monofun_cfun_arg])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   150
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   151
text {* ch2ch - rules for the type @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   152
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   153
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   154
by (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   155
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   156
lemma ch2ch_Rep_CFunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   157
by (rule monofun_Rep_CFun2 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   158
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   159
lemma ch2ch_Rep_CFunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   160
by (rule monofun_Rep_CFun1 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   161
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   162
lemma ch2ch_Rep_CFun: "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>(Y i))"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   163
apply (rule chainI)
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   164
apply (rule monofun_cfun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   165
apply (erule chainE)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   166
apply (erule chainE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   167
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   168
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   169
text {* contlub, cont properties of @{term Rep_CFun} in both arguments *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   170
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   171
lemma contlub_cfun: 
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   172
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. F i\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   173
apply (simp only: contlub_cfun_fun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   174
apply (simp only: contlub_cfun_arg)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   175
apply (rule diag_lub)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   176
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   177
apply (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   178
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   179
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   180
lemma cont_cfun: 
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   181
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> range (\<lambda>i. F i\<cdot>(Y i)) <<| (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   182
apply (rule thelubE)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   183
apply (simp only: ch2ch_Rep_CFun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   184
apply (simp only: contlub_cfun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   185
done
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   186
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   187
text {* strictness *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   188
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   189
lemma strictI: "f\<cdot>x = \<bottom> \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   190
apply (rule UU_I)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   191
apply (erule subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   192
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   193
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   194
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   195
text {* the lub of a chain of continous functions is monotone *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   196
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   197
lemma lub_cfun_mono: "chain F \<Longrightarrow> monofun (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   198
apply (drule ch2ch_monofun [OF monofun_Rep_CFun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   199
apply (simp add: thelub_fun [symmetric])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   200
apply (erule monofun_lub_fun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   201
apply (simp add: monofun_Rep_CFun2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   202
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   203
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   204
text {* a lemma about the exchange of lubs for type @{typ "'a -> 'b"}: uses MF2 lemmas from Cont.thy *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   205
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   206
lemma ex_lub_cfun: "[| chain(F); chain(Y) |] ==> 
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   207
                lub(range(%j. lub(range(%i. F(j)$(Y i))))) = 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   208
                lub(range(%i. lub(range(%j. F(j)$(Y i)))))"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   209
by (simp add: diag_lub ch2ch_Rep_CFunL ch2ch_Rep_CFunR)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   210
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   211
text {* the lub of a chain of cont. functions is continuous *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   212
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   213
lemma cont_lub_cfun: "chain F \<Longrightarrow> cont (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   214
apply (rule cont2cont_lub)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   215
apply (erule monofun_Rep_CFun [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   216
apply (rule cont_Rep_CFun2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   217
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   218
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   219
text {* type @{typ "'a -> 'b"} is chain complete *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   220
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   221
lemma lub_cfun: "chain F \<Longrightarrow> range F <<| (LAM x. LUB i. F i$x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   222
apply (subst thelub_fun [symmetric])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   223
apply (erule monofun_Rep_CFun [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   224
apply (erule typedef_is_lub [OF type_definition_CFun less_cfun_def adm_CFun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   225
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   226
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   227
lemmas thelub_cfun = lub_cfun [THEN thelubI, standard]
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   228
 -- {* @{thm thelub_cfun} *} (* chain F \<Longrightarrow> lub (range F) = (\<Lambda>x. \<Squnion>i. F i\<cdot>x) *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   229
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   230
subsection {* Miscellaneous *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   231
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   232
text {* Monotonicity of @{term Abs_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   233
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   234
lemma semi_monofun_Abs_CFun: "[| cont(f); cont(g); f<<g|] ==> Abs_CFun(f)<<Abs_CFun(g)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   235
by (simp add: less_cfun_def Abs_CFun_inverse2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   236
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   237
text {* for compatibility with old HOLCF-Version *}
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   238
lemma inst_cfun_pcpo: "\<bottom> = (\<Lambda> x. \<bottom>)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   239
by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   240
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   241
subsection {* Continuity of application *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   242
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   243
text {* cont2cont lemma for @{term Rep_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   244
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   245
lemma cont2cont_Rep_CFun:
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   246
  "\<lbrakk>cont f; cont t\<rbrakk> \<Longrightarrow> cont (\<lambda>x. (f x)\<cdot>(t x))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   247
by (best intro: cont2cont_app2 cont_const cont_Rep_CFun cont_Rep_CFun2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   248
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   249
text {* cont2mono Lemma for @{term "%x. LAM y. c1(x)(y)"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   250
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   251
lemma cont2mono_LAM:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   252
assumes p1: "!!x. cont(c1 x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   253
assumes p2: "!!y. monofun(%x. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   254
shows "monofun(%x. LAM y. c1 x y)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   255
apply (rule monofunI)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   256
apply (rule less_cfun_ext)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   257
apply (simp add: p1)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   258
apply (erule p2 [THEN monofunE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   259
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   260
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   261
text {* cont2cont Lemma for @{term "%x. LAM y. c1 x y"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   262
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   263
lemma cont2cont_LAM:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   264
assumes p1: "!!x. cont(c1 x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   265
assumes p2: "!!y. cont(%x. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   266
shows "cont(%x. LAM y. c1 x y)"
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   267
apply (rule cont_Abs_CFun)
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   268
apply (simp add: p1 CFun_def)
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   269
apply (simp add: p2 cont2cont_CF1L_rev)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   270
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   271
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   272
text {* cont2cont tactic *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   273
16055
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   274
lemmas cont_lemmas1 =
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   275
  cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   276
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   277
text {*
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   278
  Continuity simproc by Brian Huffman.
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   279
  Given the term @{term "cont f"}, the procedure tries to
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   280
  construct the theorem @{prop "cont f == True"}. If this
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   281
  theorem cannot be completely solved by the introduction
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   282
  rules, then the procedure returns a conditional rewrite
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   283
  rule with the unsolved subgoals as premises.
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   284
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   285
16055
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   286
ML_setup {*
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   287
local
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   288
  val rules = thms "cont_lemmas1";
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   289
  fun solve_cont sg _ t =
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   290
    let val tr = instantiate' [] [SOME (cterm_of sg t)] Eq_TrueI;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   291
        val tac = REPEAT_ALL_NEW (resolve_tac rules) 1;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   292
    in Option.map fst (Seq.pull (tac tr)) end;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   293
in
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   294
  val cont_proc = Simplifier.simproc (Theory.sign_of (the_context ()))
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   295
        "continuity" ["cont f"] solve_cont;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   296
end;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   297
Addsimprocs [cont_proc];
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   298
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   299
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   300
(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   301
(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   302
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   303
text {* function application is strict in its first argument *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   304
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   305
lemma Rep_CFun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   306
by (simp add: Rep_CFun_strict)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   307
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   308
text {* some lemmata for functions with flat/chfin domain/range types *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   309
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   310
lemma chfin_Rep_CFunR: "chain (Y::nat => 'a::cpo->'b::chfin)  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   311
      ==> !s. ? n. lub(range(Y))$s = Y n$s"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   312
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   313
apply (subst contlub_cfun_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   314
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   315
apply (fast intro!: thelubI chfin lub_finch2 chfin2finch ch2ch_Rep_CFunL)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   316
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   317
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   318
subsection {* Continuous injection-retraction pairs *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   319
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   320
text {* Continuous retractions are strict. *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   321
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   322
lemma retraction_strict:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   323
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   324
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   325
apply (drule_tac x="\<bottom>" in spec)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   326
apply (erule subst)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   327
apply (rule monofun_cfun_arg)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   328
apply (rule minimal)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   329
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   330
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   331
lemma injection_eq:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   332
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x = g\<cdot>y) = (x = y)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   333
apply (rule iffI)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   334
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   335
apply simp
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   336
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   337
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   338
16314
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   339
lemma injection_less:
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   340
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x \<sqsubseteq> g\<cdot>y) = (x \<sqsubseteq> y)"
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   341
apply (rule iffI)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   342
apply (drule_tac f=f in monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   343
apply simp
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   344
apply (erule monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   345
done
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   346
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   347
lemma injection_defined_rev:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   348
  "\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; g\<cdot>z = \<bottom>\<rbrakk> \<Longrightarrow> z = \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   349
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   350
apply (simp add: retraction_strict)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   351
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   352
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   353
lemma injection_defined:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   354
  "\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; z \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> g\<cdot>z \<noteq> \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   355
by (erule contrapos_nn, rule injection_defined_rev)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   356
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   357
text {* propagation of flatness and chain-finiteness by retractions *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   358
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   359
lemma chfin2chfin:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   360
  "\<forall>y. (f::'a::chfin \<rightarrow> 'b)\<cdot>(g\<cdot>y) = y
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   361
    \<Longrightarrow> \<forall>Y::nat \<Rightarrow> 'b. chain Y \<longrightarrow> (\<exists>n. max_in_chain n Y)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   362
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   363
apply (drule_tac f=g in chain_monofun)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   364
apply (drule chfin [rule_format])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   365
apply (unfold max_in_chain_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   366
apply (simp add: injection_eq)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   367
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   368
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   369
lemma flat2flat:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   370
  "\<forall>y. (f::'a::flat \<rightarrow> 'b::pcpo)\<cdot>(g\<cdot>y) = y
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   371
    \<Longrightarrow> \<forall>x y::'b. x \<sqsubseteq> y \<longrightarrow> x = \<bottom> \<or> x = y"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   372
apply clarify
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   373
apply (drule_tac f=g in monofun_cfun_arg)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   374
apply (drule ax_flat [rule_format])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   375
apply (erule disjE)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   376
apply (simp add: injection_defined_rev)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   377
apply (simp add: injection_eq)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   378
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   379
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   380
text {* a result about functions with flat codomain *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   381
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   382
lemma flat_eqI: "\<lbrakk>(x::'a::flat) \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> x = y"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   383
by (drule ax_flat [rule_format], simp)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   384
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   385
lemma flat_codom:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   386
  "f\<cdot>x = (c::'b::flat) \<Longrightarrow> f\<cdot>\<bottom> = \<bottom> \<or> (\<forall>z. f\<cdot>z = c)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   387
apply (case_tac "f\<cdot>x = \<bottom>")
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   388
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   389
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   390
apply (erule_tac t="\<bottom>" in subst)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   391
apply (rule minimal [THEN monofun_cfun_arg])
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   392
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   393
apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   394
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   395
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   396
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   397
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   398
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   399
subsection {* Identity and composition *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   400
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   401
consts
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   402
  ID      :: "'a \<rightarrow> 'a"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   403
  cfcomp  :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   404
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   405
syntax  "@oo" :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c" (infixr "oo" 100)
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   406
     
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   407
translations  "f1 oo f2" == "cfcomp$f1$f2"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   408
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   409
defs
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   410
  ID_def: "ID \<equiv> (\<Lambda> x. x)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   411
  oo_def: "cfcomp \<equiv> (\<Lambda> f g x. f\<cdot>(g\<cdot>x))" 
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   412
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   413
lemma ID1 [simp]: "ID\<cdot>x = x"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   414
by (simp add: ID_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   415
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   416
lemma cfcomp1: "(f oo g) = (\<Lambda> x. f\<cdot>(g\<cdot>x))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   417
by (simp add: oo_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   418
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   419
lemma cfcomp2 [simp]: "(f oo g)\<cdot>x = f\<cdot>(g\<cdot>x)"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   420
by (simp add: cfcomp1)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   421
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   422
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   423
  Show that interpretation of (pcpo,@{text "_->_"}) is a category.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   424
  The class of objects is interpretation of syntactical class pcpo.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   425
  The class of arrows  between objects @{typ 'a} and @{typ 'b} is interpret. of @{typ "'a -> 'b"}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   426
  The identity arrow is interpretation of @{term ID}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   427
  The composition of f and g is interpretation of @{text "oo"}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   428
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   429
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   430
lemma ID2 [simp]: "f oo ID = f"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   431
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   432
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   433
lemma ID3 [simp]: "ID oo f = f"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   434
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   435
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   436
lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   437
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   438
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   439
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   440
subsection {* Strictified functions *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   441
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   442
defaultsort pcpo
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   443
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   444
consts  
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   445
  Istrictify :: "('a \<rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   446
  strictify  :: "('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'b"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   447
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   448
defs
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   449
  Istrictify_def: "Istrictify f x \<equiv> if x = \<bottom> then \<bottom> else f\<cdot>x"    
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   450
  strictify_def:  "strictify \<equiv> (\<Lambda> f x. Istrictify f x)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   451
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   452
text {* results about strictify *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   453
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   454
lemma Istrictify1: "Istrictify f \<bottom> = \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   455
by (simp add: Istrictify_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   456
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   457
lemma Istrictify2: "x \<noteq> \<bottom> \<Longrightarrow> Istrictify f x = f\<cdot>x"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   458
by (simp add: Istrictify_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   459
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   460
lemma cont_Istrictify1: "cont (\<lambda>f. Istrictify f x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   461
apply (case_tac "x = \<bottom>")
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   462
apply (simp add: Istrictify1)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   463
apply (simp add: Istrictify2)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   464
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   465
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   466
lemma monofun_Istrictify2: "monofun (\<lambda>x. Istrictify f x)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   467
apply (rule monofunI)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   468
apply (simp add: Istrictify_def monofun_cfun_arg)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   469
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   470
apply (simp add: eq_UU_iff)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   471
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   472
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   473
lemma contlub_Istrictify2: "contlub (\<lambda>x. Istrictify f x)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   474
apply (rule contlubI)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   475
apply (case_tac "lub (range Y) = \<bottom>")
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   476
apply (simp add: Istrictify1 chain_UU_I thelub_const)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   477
apply (simp add: Istrictify2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   478
apply (simp add: contlub_cfun_arg)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   479
apply (rule lub_equal2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   480
apply (rule chain_mono2 [THEN exE])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   481
apply (erule chain_UU_I_inverse2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   482
apply (assumption)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   483
apply (blast intro: Istrictify2 [symmetric])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   484
apply (erule chain_monofun)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   485
apply (erule monofun_Istrictify2 [THEN ch2ch_monofun])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   486
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   487
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   488
lemmas cont_Istrictify2 =
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   489
  monocontlub2cont [OF monofun_Istrictify2 contlub_Istrictify2, standard]
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   490
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   491
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   492
apply (unfold strictify_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   493
apply (simp add: cont_Istrictify1 cont_Istrictify2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   494
apply (rule Istrictify1)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   495
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   496
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   497
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   498
apply (unfold strictify_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   499
apply (simp add: cont_Istrictify1 cont_Istrictify2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   500
apply (erule Istrictify2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   501
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   502
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   503
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   504
by simp
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   505
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   506
end