src/HOLCF/Cfun.thy
author huffman
Mon, 23 May 2005 23:32:07 +0200
changeset 16055 58186c507750
parent 15641 b389f108c485
child 16070 4a83dd540b88
permissions -rw-r--r--
moved continuity simproc to Cfun.thy
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
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     5
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     6
Definition of the type ->  of continuous functions.
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
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     9
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    10
header {* The type of continuous functions *}
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    11
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    12
theory Cfun
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    13
imports Cont
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    14
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    15
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    16
defaultsort cpo
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    17
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    18
subsection {* Definition of continuous function type *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    19
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    20
typedef (CFun)  ('a, 'b) "->" (infixr 0) = "{f::'a => 'b. cont f}"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    21
by (rule exI, rule CfunI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    22
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    23
syntax
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    24
	Rep_CFun  :: "('a -> 'b) => ('a => 'b)" ("_$_" [999,1000] 999)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    25
                                                (* application      *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    26
        Abs_CFun  :: "('a => 'b) => ('a -> 'b)" (binder "LAM " 10)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    27
                                                (* abstraction      *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    28
        less_cfun :: "[('a -> 'b),('a -> 'b)]=>bool"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    29
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    30
syntax (xsymbols)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    31
  "->"		:: "[type, type] => type"      ("(_ \<rightarrow>/ _)" [1,0]0)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    32
  "LAM "	:: "[idts, 'a => 'b] => ('a -> 'b)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    33
					("(3\<Lambda>_./ _)" [0, 10] 10)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    34
  Rep_CFun      :: "('a -> 'b) => ('a => 'b)"  ("(_\<cdot>_)" [999,1000] 999)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    35
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    36
syntax (HTML output)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    37
  Rep_CFun      :: "('a -> 'b) => ('a => 'b)"  ("(_\<cdot>_)" [999,1000] 999)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    38
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    39
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    40
  Derive old type definition rules for @{term Abs_CFun} \& @{term Rep_CFun}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    41
  @{term Rep_CFun} and @{term Abs_CFun} should be replaced by
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    42
  @{term Rep_Cfun} and @{term Abs_Cfun} in future.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    43
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    44
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    45
lemma Rep_Cfun: "Rep_CFun fo : CFun"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    46
by (rule Rep_CFun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    47
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    48
lemma Rep_Cfun_inverse: "Abs_CFun (Rep_CFun fo) = fo"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    49
by (rule Rep_CFun_inverse)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    50
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    51
lemma Abs_Cfun_inverse: "f:CFun==>Rep_CFun(Abs_CFun f)=f"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    52
by (erule Abs_CFun_inverse)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    53
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    54
text {* Additional lemma about the isomorphism between
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    55
        @{typ "'a -> 'b"} and @{term Cfun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    56
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    57
lemma Abs_Cfun_inverse2: "cont f ==> Rep_CFun (Abs_CFun f) = f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    58
apply (rule Abs_Cfun_inverse)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    59
apply (unfold CFun_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    60
apply (erule mem_Collect_eq [THEN ssubst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    61
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    62
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    63
text {* Simplification of application *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    64
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    65
lemma Cfunapp2: "cont f ==> (Abs_CFun f)$x = f x"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    66
by (erule Abs_Cfun_inverse2 [THEN fun_cong])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    67
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    68
text {* Beta - equality for continuous functions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    69
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    70
lemma beta_cfun: "cont(c1) ==> (LAM x .c1 x)$u = c1 u"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    71
by (rule Cfunapp2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    72
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    73
text {* Eta - equality for continuous functions *}
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    74
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    75
lemma eta_cfun: "(LAM x. f$x) = f"
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    76
by (rule Rep_CFun_inverse)
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    77
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    78
subsection {* Type @{typ "'a -> 'b"} is a partial order *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    79
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    80
instance "->"  :: (cpo, cpo) sq_ord ..
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    81
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    82
defs (overloaded)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    83
  less_cfun_def: "(op <<) == (% fo1 fo2. Rep_CFun fo1 << Rep_CFun fo2 )"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    84
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    85
lemma refl_less_cfun: "(f::'a->'b) << f"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    86
by (unfold less_cfun_def, rule refl_less)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    87
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    88
lemma antisym_less_cfun: 
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    89
        "[|(f1::'a->'b) << f2; f2 << f1|] ==> f1 = f2"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    90
by (unfold less_cfun_def, rule Rep_CFun_inject[THEN iffD1], rule antisym_less)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    91
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    92
lemma trans_less_cfun: 
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    93
        "[|(f1::'a->'b) << f2; f2 << f3|] ==> f1 << f3"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    94
by (unfold less_cfun_def, rule trans_less)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    95
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    96
instance "->" :: (cpo, cpo) po
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    97
by intro_classes
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    98
  (assumption | rule refl_less_cfun antisym_less_cfun trans_less_cfun)+
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    99
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   100
text {* for compatibility with old HOLCF-Version *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   101
lemma inst_cfun_po: "(op <<)=(%f1 f2. Rep_CFun f1 << Rep_CFun f2)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   102
apply (fold less_cfun_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   103
apply (rule refl)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   104
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   105
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   106
text {* lemmas about application of continuous functions *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   107
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   108
lemma cfun_cong: "[| f=g; x=y |] ==> f$x = g$y"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   109
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   110
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   111
lemma cfun_fun_cong: "f=g ==> f$x = g$x"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   112
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   113
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   114
lemma cfun_arg_cong: "x=y ==> f$x = f$y"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   115
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   116
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   117
text {* access to @{term less_cfun} in class po *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   118
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   119
lemma less_cfun: "( f1 << f2 ) = (Rep_CFun(f1) << Rep_CFun(f2))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   120
by (simp add: inst_cfun_po)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   121
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   122
subsection {* Type @{typ "'a -> 'b"} is pointed *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   123
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   124
lemma minimal_cfun: "Abs_CFun(% x. UU) << f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   125
apply (subst less_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   126
apply (subst Abs_Cfun_inverse2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   127
apply (rule cont_const)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   128
apply (rule minimal_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   129
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   130
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   131
lemmas UU_cfun_def = minimal_cfun [THEN minimal2UU, symmetric, standard]
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   132
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   133
lemma least_cfun: "? x::'a->'b::pcpo.!y. x<<y"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   134
apply (rule_tac x = "Abs_CFun (% x. UU) " in exI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   135
apply (rule minimal_cfun [THEN allI])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   136
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   137
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   138
subsection {* Monotonicity of application *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   139
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   140
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   141
  @{term Rep_CFun} yields continuous functions in @{typ "'a => 'b"}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   142
  This is continuity of @{term Rep_CFun} in its 'second' argument:
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   143
  @{prop "cont_Rep_CFun2 ==> monofun_Rep_CFun2 & contlub_Rep_CFun2"}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   144
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   145
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   146
lemma cont_Rep_CFun2: "cont(Rep_CFun(fo))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   147
apply (rule_tac P = "cont" in CollectD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   148
apply (fold CFun_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   149
apply (rule Rep_Cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   150
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   151
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   152
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   153
 -- {* @{thm monofun_Rep_CFun2} *} (* monofun(Rep_CFun(?fo)) *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   154
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   155
lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   156
 -- {* @{thm contlub_Rep_CFun2} *} (* contlub(Rep_CFun(?fo)) *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   157
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   158
text {* expanded thms @{thm [source] cont_Rep_CFun2}, @{thm [source] contlub_Rep_CFun2} look nice with mixfix syntax *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   159
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   160
lemmas cont_cfun_arg = cont_Rep_CFun2 [THEN contE, THEN spec, THEN mp]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   161
  -- {* @{thm cont_cfun_arg} *} (* chain(x1) ==> range (%i. fo3$(x1 i)) <<| fo3$(lub (range ?x1))    *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   162
 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   163
lemmas contlub_cfun_arg = contlub_Rep_CFun2 [THEN contlubE, THEN spec, THEN mp]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   164
 -- {* @{thm contlub_cfun_arg} *} (* chain(?x1) ==> ?fo4$(lub (range ?x1)) = lub (range (%i. ?fo4$(?x1 i))) *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   165
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   166
text {* @{term Rep_CFun} is monotone in its 'first' argument *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   167
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   168
lemma monofun_Rep_CFun1: "monofun(Rep_CFun)"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   169
apply (rule monofunI [rule_format])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   170
apply (erule less_cfun [THEN subst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   171
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   172
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   173
text {* monotonicity of application @{term Rep_CFun} in mixfix syntax @{text "[_]_"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   174
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   175
lemma monofun_cfun_fun: "f1 << f2 ==> f1$x << f2$x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   176
apply (rule_tac x = "x" in spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   177
apply (rule less_fun [THEN subst])
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   178
apply (erule monofun_Rep_CFun1 [THEN monofunE [rule_format]])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   179
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   180
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   181
lemmas monofun_cfun_arg = monofun_Rep_CFun2 [THEN monofunE [rule_format], standard]
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   182
 -- {* @{thm monofun_cfun_arg} *} (* ?x2 << ?x1 ==> ?fo5$?x2 << ?fo5$?x1 *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   183
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   184
lemma chain_monofun: "chain Y ==> chain (%i. f\<cdot>(Y i))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   185
apply (rule chainI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   186
apply (rule monofun_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   187
apply (erule chainE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   188
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   189
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   190
text {* monotonicity of @{term Rep_CFun} in both arguments in mixfix syntax @{text "[_]_"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   191
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   192
lemma monofun_cfun: "[|f1<<f2;x1<<x2|] ==> f1$x1 << f2$x2"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   193
apply (rule trans_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   194
apply (erule monofun_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   195
apply (erule monofun_cfun_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   196
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   197
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   198
lemma strictI: "f$x = UU ==> f$UU = UU"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   199
apply (rule eq_UU_iff [THEN iffD2])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   200
apply (erule subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   201
apply (rule minimal [THEN monofun_cfun_arg])
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
subsection {* Type @{typ "'a -> 'b"} is a cpo *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   205
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   206
text {* ch2ch - rules for the type @{typ "'a -> 'b"} use MF2 lemmas from Cont.thy *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   207
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   208
lemma ch2ch_Rep_CFunR: "chain(Y) ==> chain(%i. f$(Y i))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   209
by (erule monofun_Rep_CFun2 [THEN ch2ch_MF2R])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   210
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   211
lemmas ch2ch_Rep_CFunL = monofun_Rep_CFun1 [THEN ch2ch_MF2L, standard]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   212
 -- {* @{thm ch2ch_Rep_CFunL} *} (* chain(?F) ==> chain (%i. ?F i$?x) *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   213
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   214
text {* the lub of a chain of continous functions is monotone: uses MF2 lemmas from Cont.thy *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   215
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   216
lemma lub_cfun_mono: "chain(F) ==> monofun(% x. lub(range(% j.(F j)$x)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   217
apply (rule lub_MF2_mono)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   218
apply (rule monofun_Rep_CFun1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   219
apply (rule monofun_Rep_CFun2 [THEN allI])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   220
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   221
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   222
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   223
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
   224
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   225
lemma ex_lubcfun: "[| chain(F); chain(Y) |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   226
                lub(range(%j. lub(range(%i. F(j)$(Y i))))) = 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   227
                lub(range(%i. lub(range(%j. F(j)$(Y i)))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   228
apply (rule ex_lubMF2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   229
apply (rule monofun_Rep_CFun1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   230
apply (rule monofun_Rep_CFun2 [THEN allI])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   231
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   232
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   233
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   234
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   235
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
   236
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   237
lemma cont_lubcfun: "chain(F) ==> cont(% x. lub(range(% j. F(j)$x)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   238
apply (rule monocontlub2cont)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   239
apply (erule lub_cfun_mono)
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   240
apply (rule contlubI [rule_format])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   241
apply (subst contlub_cfun_arg [THEN ext])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   242
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   243
apply (erule ex_lubcfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   244
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   245
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   246
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   247
text {* type @{typ "'a -> 'b"} is chain complete *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   248
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   249
lemma lub_cfun: "chain(CCF) ==> range(CCF) <<| (LAM x. lub(range(% i. CCF(i)$x)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   250
apply (rule is_lubI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   251
apply (rule ub_rangeI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   252
apply (subst less_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   253
apply (subst Abs_Cfun_inverse2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   254
apply (erule cont_lubcfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   255
apply (rule lub_fun [THEN is_lubD1, THEN ub_rangeD])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   256
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   257
apply (subst less_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   258
apply (subst Abs_Cfun_inverse2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   259
apply (erule cont_lubcfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   260
apply (rule lub_fun [THEN is_lub_lub])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   261
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   262
apply (erule monofun_Rep_CFun1 [THEN ub2ub_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   263
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   264
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   265
lemmas thelub_cfun = lub_cfun [THEN thelubI, standard]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   266
 -- {* @{thm thelub_cfun} *} (* 
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   267
chain(?CCF1) ==>  lub (range ?CCF1) = (LAM x. lub (range (%i. ?CCF1 i$x)))
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   268
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   269
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   270
lemma cpo_cfun: "chain(CCF::nat=>('a->'b)) ==> ? x. range(CCF) <<| x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   271
apply (rule exI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   272
apply (erule lub_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   273
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   274
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   275
instance "->" :: (cpo, cpo) cpo
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   276
by intro_classes (rule cpo_cfun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   277
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   278
subsection {* Miscellaneous *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   279
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   280
text {* Extensionality in @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   281
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   282
lemma ext_cfun: "(!!x. f$x = g$x) ==> f = g"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   283
apply (rule Rep_CFun_inject [THEN iffD1])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   284
apply (rule ext)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   285
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   286
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   287
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   288
text {* Monotonicity of @{term Abs_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   289
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   290
lemma semi_monofun_Abs_CFun: "[| cont(f); cont(g); f<<g|] ==> Abs_CFun(f)<<Abs_CFun(g)"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   291
by (simp add: less_cfun Abs_Cfun_inverse2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   292
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   293
text {* Extensionality wrt. @{term "op <<"} in @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   294
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   295
lemma less_cfun2: "(!!x. f$x << g$x) ==> f << g"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   296
apply (rule_tac t = "f" in Rep_Cfun_inverse [THEN subst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   297
apply (rule_tac t = "g" in Rep_Cfun_inverse [THEN subst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   298
apply (rule semi_monofun_Abs_CFun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   299
apply (rule cont_Rep_CFun2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   300
apply (rule cont_Rep_CFun2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   301
apply (rule less_fun [THEN iffD2])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   302
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   303
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   304
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   305
subsection {* Class instance of @{typ "'a -> 'b"} for class pcpo *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   306
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   307
instance "->" :: (cpo, pcpo) pcpo
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   308
by (intro_classes, rule least_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   309
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   310
text {* for compatibility with old HOLCF-Version *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   311
lemma inst_cfun_pcpo: "UU = Abs_CFun(%x. UU)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   312
apply (simp add: UU_def UU_cfun_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   313
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   314
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   315
defaultsort pcpo
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   316
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   317
subsection {* Continuity of application *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   318
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   319
text {* the contlub property for @{term Rep_CFun} its 'first' argument *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   320
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   321
lemma contlub_Rep_CFun1: "contlub(Rep_CFun)"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   322
apply (rule contlubI [rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   323
apply (rule ext)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   324
apply (subst thelub_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   325
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   326
apply (subst Cfunapp2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   327
apply (erule cont_lubcfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   328
apply (subst thelub_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   329
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   330
apply (rule refl)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   331
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   332
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   333
text {* the cont property for @{term Rep_CFun} in its first argument *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   334
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   335
lemma cont_Rep_CFun1: "cont(Rep_CFun)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   336
apply (rule monocontlub2cont)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   337
apply (rule monofun_Rep_CFun1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   338
apply (rule contlub_Rep_CFun1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   339
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   340
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   341
text {* contlub, cont properties of @{term Rep_CFun} in its first argument in mixfix @{text "_[_]"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   342
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   343
lemma contlub_cfun_fun: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   344
"chain(FY) ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   345
  lub(range FY)$x = lub(range (%i. FY(i)$x))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   346
apply (rule trans)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   347
apply (erule contlub_Rep_CFun1 [THEN contlubE, THEN spec, THEN mp, THEN fun_cong])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   348
apply (subst thelub_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   349
apply (erule monofun_Rep_CFun1 [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   350
apply (rule refl)
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
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   353
lemma cont_cfun_fun: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   354
"chain(FY) ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   355
  range(%i. FY(i)$x) <<| lub(range FY)$x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   356
apply (rule thelubE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   357
apply (erule ch2ch_Rep_CFunL)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   358
apply (erule contlub_cfun_fun [symmetric])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   359
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   360
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   361
text {* contlub, cont  properties of @{term Rep_CFun} in both argument in mixfix @{text "_[_]"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   362
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   363
lemma contlub_cfun: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   364
"[|chain(FY);chain(TY)|] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   365
  (lub(range FY))$(lub(range TY)) = lub(range(%i. FY(i)$(TY i)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   366
apply (rule contlub_CF2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   367
apply (rule cont_Rep_CFun1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   368
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   369
apply (rule cont_Rep_CFun2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   370
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   371
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   372
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   373
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   374
lemma cont_cfun: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   375
"[|chain(FY);chain(TY)|] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   376
  range(%i.(FY i)$(TY i)) <<| (lub (range FY))$(lub(range TY))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   377
apply (rule thelubE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   378
apply (rule monofun_Rep_CFun1 [THEN ch2ch_MF2LR])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   379
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   380
apply (rule monofun_Rep_CFun2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   381
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   382
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   383
apply (erule contlub_cfun [symmetric])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   384
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   385
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   386
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   387
text {* cont2cont lemma for @{term Rep_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   388
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   389
lemma cont2cont_Rep_CFun: "[|cont(%x. ft x);cont(%x. tt x)|] ==> cont(%x. (ft x)$(tt x))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   390
apply (best intro: cont2cont_app2 cont_const cont_Rep_CFun1 cont_Rep_CFun2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   391
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   392
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   393
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
   394
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   395
lemma cont2mono_LAM:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   396
assumes p1: "!!x. cont(c1 x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   397
assumes p2: "!!y. monofun(%x. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   398
shows "monofun(%x. LAM y. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   399
apply (rule monofunI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   400
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   401
apply (subst less_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   402
apply (subst less_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   403
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   404
apply (subst beta_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   405
apply (rule p1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   406
apply (subst beta_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   407
apply (rule p1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   408
apply (erule p2 [THEN monofunE, THEN spec, THEN spec, THEN mp])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   409
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   410
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   411
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
   412
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   413
lemma cont2cont_LAM:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   414
assumes p1: "!!x. cont(c1 x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   415
assumes p2: "!!y. cont(%x. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   416
shows "cont(%x. LAM y. c1 x y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   417
apply (rule monocontlub2cont)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   418
apply (rule p1 [THEN cont2mono_LAM])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   419
apply (rule p2 [THEN cont2mono])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   420
apply (rule contlubI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   421
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   422
apply (subst thelub_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   423
apply (rule p1 [THEN cont2mono_LAM, THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   424
apply (rule p2 [THEN cont2mono])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   425
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   426
apply (rule_tac f = "Abs_CFun" in arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   427
apply (rule ext)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   428
apply (subst p1 [THEN beta_cfun, THEN ext])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   429
apply (erule p2 [THEN cont2contlub, THEN contlubE, THEN spec, THEN mp])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   430
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   431
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
   432
text {* cont2cont Lemma for @{term "%x. LAM y. c1 x$y"} *}
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
   433
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
   434
lemma cont2cont_eta: "cont c1 ==> cont (%x. LAM y. c1 x$y)"
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
   435
by (simp only: eta_cfun)
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
   436
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   437
text {* cont2cont tactic *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   438
16055
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   439
lemmas cont_lemmas1 =
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   440
  cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   441
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   442
text {*
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   443
  Continuity simproc by Brian Huffman.
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   444
  Given the term @{term "cont f"}, the procedure tries to
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   445
  construct the theorem @{prop "cont f == True"}. If this
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   446
  theorem cannot be completely solved by the introduction
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   447
  rules, then the procedure returns a conditional rewrite
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   448
  rule with the unsolved subgoals as premises.
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   449
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   450
16055
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   451
ML_setup {*
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   452
local
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   453
  val rules = thms "cont_lemmas1";
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   454
  fun solve_cont sg _ t =
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   455
    let val tr = instantiate' [] [SOME (cterm_of sg t)] Eq_TrueI;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   456
        val tac = REPEAT_ALL_NEW (resolve_tac rules) 1;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   457
    in Option.map fst (Seq.pull (tac tr)) end;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   458
in
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   459
  val cont_proc = Simplifier.simproc (Theory.sign_of (the_context ()))
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   460
        "continuity" ["cont f"] solve_cont;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   461
end;
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   462
Addsimprocs [cont_proc];
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   463
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   464
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   465
text {* HINT: @{text cont_tac} is now installed in simplifier in Lift.ML ! *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   466
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   467
(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   468
(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   469
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   470
text {* function application @{text "_[_]"} is strict in its first arguments *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   471
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   472
lemma strict_Rep_CFun1 [simp]: "(UU::'a::cpo->'b)$x = (UU::'b)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   473
by (simp add: inst_cfun_pcpo beta_cfun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   474
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   475
text {* Instantiate the simplifier *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   476
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   477
declare beta_cfun [simp]
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   478
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   479
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
   480
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   481
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
   482
      ==> !s. ? n. lub(range(Y))$s = Y n$s"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   483
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   484
apply (subst contlub_cfun_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   485
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   486
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
   487
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   488
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   489
subsection {* Continuous isomorphisms *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   490
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   491
text {* Continuous isomorphisms are strict. A proof for embedding projection pairs is similar. *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   492
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   493
lemma iso_strict: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   494
"!!f g.[|!y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a) |]  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   495
  ==> f$UU=UU & g$UU=UU"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   496
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   497
apply (rule UU_I)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   498
apply (rule_tac s = "f$ (g$ (UU::'b))" and t = "UU::'b" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   499
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   500
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   501
apply (rule UU_I)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   502
apply (rule_tac s = "g$ (f$ (UU::'a))" and t = "UU::'a" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   503
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   504
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   505
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   506
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   507
lemma isorep_defined: "[|!x. rep$(ab$x)=x;!y. ab$(rep$y)=y; z~=UU|] ==> rep$z ~= UU"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   508
apply (erule contrapos_nn)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   509
apply (drule_tac f = "ab" in cfun_arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   510
apply (erule box_equals)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   511
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   512
apply (erule iso_strict [THEN conjunct1])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   513
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   514
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   515
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   516
lemma isoabs_defined: "[|!x. rep$(ab$x) = x;!y. ab$(rep$y)=y ; z~=UU|] ==> ab$z ~= UU"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   517
apply (erule contrapos_nn)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   518
apply (drule_tac f = "rep" in cfun_arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   519
apply (erule box_equals)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   520
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   521
apply (erule iso_strict [THEN conjunct2])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   522
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   523
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   524
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   525
text {* propagation of flatness and chain-finiteness by continuous isomorphisms *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   526
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   527
lemma chfin2chfin: "!!f g.[|! Y::nat=>'a. chain Y --> (? n. max_in_chain n Y);
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   528
  !y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a::chfin) |]  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   529
  ==> ! Y::nat=>'b. chain Y --> (? n. max_in_chain n Y)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   530
apply (unfold max_in_chain_def)
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   531
apply (clarify)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   532
apply (rule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   533
apply (rule_tac P = "chain (%i. g$ (Y i))" in mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   534
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   535
apply (erule ch2ch_Rep_CFunR)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   536
apply (rule exI)
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   537
apply (clarify)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   538
apply (rule_tac s = "f$ (g$ (Y x))" and t = "Y (x) " in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   539
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   540
apply (rule_tac s = "f$ (g$ (Y j))" and t = "Y (j) " in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   541
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   542
apply (rule cfun_arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   543
apply (rule mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   544
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   545
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   546
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   547
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   548
lemma flat2flat: "!!f g.[|!x y::'a. x<<y --> x=UU | x=y;  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   549
  !y. f$(g$y)=(y::'b); !x. g$(f$x)=(x::'a)|] ==> !x y::'b. x<<y --> x=UU | x=y"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   550
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   551
apply (rule disjE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   552
apply (rule_tac P = "g$x<<g$y" in mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   553
apply (erule_tac [2] monofun_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   554
apply (drule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   555
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   556
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   557
apply (rule trans)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   558
apply (rule_tac s = "f$ (g$x) " and t = "x" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   559
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   560
apply (erule cfun_arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   561
apply (rule iso_strict [THEN conjunct1])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   562
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   563
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   564
apply (rule disjI2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   565
apply (rule_tac s = "f$ (g$x) " and t = "x" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   566
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   567
apply (rule_tac s = "f$ (g$y) " and t = "y" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   568
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   569
apply (erule cfun_arg_cong)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   570
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   571
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   572
text {* a result about functions with flat codomain *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   573
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   574
lemma flat_codom: "f$(x::'a)=(c::'b::flat) ==> f$(UU::'a)=(UU::'b) | (!z. f$(z::'a)=c)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   575
apply (case_tac "f$ (x::'a) = (UU::'b) ")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   576
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   577
apply (rule UU_I)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   578
apply (rule_tac s = "f$ (x) " and t = "UU::'b" in subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   579
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   580
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   581
apply (case_tac "f$ (UU::'a) = (UU::'b) ")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   582
apply (erule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   583
apply (rule disjI2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   584
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   585
apply (erule subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   586
apply (rule_tac a = "f$ (UU::'a) " in refl [THEN box_equals])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   587
apply (rule_tac fo5 = "f" in minimal [THEN monofun_cfun_arg, THEN ax_flat [THEN spec, THEN spec, THEN mp], THEN disjE])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   588
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   589
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   590
apply (rule_tac fo5 = "f" in minimal [THEN monofun_cfun_arg, THEN ax_flat [THEN spec, THEN spec, THEN mp], THEN disjE])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   591
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   592
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   593
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   594
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   595
subsection {* Strictified functions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   596
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   597
consts  
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   598
        Istrictify   :: "('a->'b)=>'a=>'b"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   599
        strictify    :: "('a->'b)->'a->'b"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   600
defs
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   601
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   602
Istrictify_def:  "Istrictify f x == if x=UU then UU else f$x"    
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   603
strictify_def:   "strictify == (LAM f x. Istrictify f x)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   604
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   605
text {* results about strictify *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   606
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   607
lemma Istrictify1: 
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   608
        "Istrictify(f)(UU)= (UU)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   609
apply (unfold Istrictify_def)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   610
apply (simp (no_asm))
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   611
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   612
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   613
lemma Istrictify2: 
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   614
        "~x=UU ==> Istrictify(f)(x)=f$x"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   615
by (simp add: Istrictify_def)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   616
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   617
lemma monofun_Istrictify1: "monofun(Istrictify)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   618
apply (rule monofunI [rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   619
apply (rule less_fun [THEN iffD2, rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   620
apply (case_tac "xa=UU")
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   621
apply (simp add: Istrictify1)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   622
apply (simp add: Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   623
apply (erule monofun_cfun_fun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   624
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   625
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   626
lemma monofun_Istrictify2: "monofun(Istrictify(f))"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   627
apply (rule monofunI [rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   628
apply (case_tac "x=UU")
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   629
apply (simp add: Istrictify1)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   630
apply (frule notUU_I)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   631
apply assumption
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   632
apply (simp add: Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   633
apply (erule monofun_cfun_arg)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   634
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   635
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   636
lemma contlub_Istrictify1: "contlub(Istrictify)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   637
apply (rule contlubI [rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   638
apply (rule ext)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   639
apply (subst thelub_fun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   640
apply (erule monofun_Istrictify1 [THEN ch2ch_monofun])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   641
apply (case_tac "x=UU")
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   642
apply (simp add: Istrictify1)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   643
apply (simp add: lub_const [THEN thelubI])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   644
apply (simp add: Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   645
apply (erule contlub_cfun_fun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   646
done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   647
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   648
lemma contlub_Istrictify2: "contlub(Istrictify(f::'a -> 'b))"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   649
apply (rule contlubI [rule_format])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   650
apply (case_tac "lub (range (Y))=UU")
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   651
apply (simp add: Istrictify1 chain_UU_I)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   652
apply (simp add: lub_const [THEN thelubI])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   653
apply (simp add: Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   654
apply (rule_tac s = "lub (range (%i. f$ (Y i)))" in trans)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   655
apply (erule contlub_cfun_arg)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   656
apply (rule lub_equal2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   657
apply (rule chain_mono2 [THEN exE])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   658
apply (erule chain_UU_I_inverse2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   659
apply (assumption)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   660
apply (blast intro: Istrictify2 [symmetric])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   661
apply (erule chain_monofun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   662
apply (erule monofun_Istrictify2 [THEN ch2ch_monofun])
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   663
done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   664
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   665
lemmas cont_Istrictify1 = contlub_Istrictify1 [THEN monofun_Istrictify1 [THEN monocontlub2cont], standard]
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   666
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   667
lemmas cont_Istrictify2 = contlub_Istrictify2 [THEN monofun_Istrictify2 [THEN monocontlub2cont], standard]
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   668
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   669
lemma strictify1 [simp]: "strictify$f$UU=UU"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   670
apply (unfold strictify_def)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   671
apply (subst beta_cfun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   672
apply (simp add: cont_Istrictify2 cont_Istrictify1 cont2cont_CF1L)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   673
apply (subst beta_cfun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   674
apply (rule cont_Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   675
apply (rule Istrictify1)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   676
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   677
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   678
lemma strictify2 [simp]: "~x=UU ==> strictify$f$x=f$x"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   679
apply (unfold strictify_def)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   680
apply (subst beta_cfun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   681
apply (simp add: cont_Istrictify2 cont_Istrictify1 cont2cont_CF1L)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   682
apply (subst beta_cfun)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   683
apply (rule cont_Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   684
apply (erule Istrictify2)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   685
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   686
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   687
subsection {* Identity and composition *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   688
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   689
consts
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   690
        ID      :: "('a::cpo) -> 'a"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   691
        cfcomp  :: "('b->'c)->(('a::cpo)->('b::cpo))->'a->('c::cpo)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   692
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   693
syntax  "@oo"   :: "('b->'c)=>('a->'b)=>'a->'c" ("_ oo _" [101,100] 100)
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   694
     
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   695
translations    "f1 oo f2" == "cfcomp$f1$f2"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   696
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   697
defs
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   698
  ID_def:        "ID ==(LAM x. x)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   699
  oo_def:        "cfcomp == (LAM f g x. f$(g$x))" 
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   700
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   701
lemma ID1 [simp]: "ID$x=x"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   702
apply (unfold ID_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   703
apply (subst beta_cfun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   704
apply (rule cont_id)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   705
apply (rule refl)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   706
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   707
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   708
lemma cfcomp1: "(f oo g)=(LAM x. f$(g$x))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   709
by (simp add: oo_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   710
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   711
lemma cfcomp2 [simp]: "(f oo g)$x=f$(g$x)"
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   712
by (simp add: cfcomp1)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   713
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   714
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   715
  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
   716
  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
   717
  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
   718
  The identity arrow is interpretation of @{term ID}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   719
  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
   720
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   721
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   722
lemma ID2 [simp]: "f oo ID = f "
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   723
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   724
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   725
lemma ID3 [simp]: "ID oo f = f "
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   726
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   727
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   728
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
   729
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   730
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   731
end