src/HOLCF/Cfun.thy
author hoelzl
Thu, 12 Nov 2009 17:21:43 +0100
changeset 33638 548a34929e98
parent 31076 99fe356cbbc2
child 35115 446c5063e4fd
permissions -rw-r--r--
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
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
    Author:     Franz Regensburger
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     3
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     4
Definition of the type ->  of continuous functions.
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
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     7
header {* The type of continuous functions *}
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     8
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
theory Cfun
29533
7f4a32134447 minimize dependencies
huffman
parents: 29530
diff changeset
    10
imports Pcpodef Ffun Product_Cpo
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    11
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    12
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    13
defaultsort cpo
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    14
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    15
subsection {* Definition of continuous function type *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    16
16699
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    17
lemma Ex_cont: "\<exists>f. cont f"
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    18
by (rule exI, rule cont_const)
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    19
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    20
lemma adm_cont: "adm cont"
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    21
by (rule admI, rule cont_lub_fun)
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
    22
17817
405fb812e738 add names to infix declarations
huffman
parents: 17816
diff changeset
    23
cpodef (CFun)  ('a, 'b) "->" (infixr "->" 0) = "{f::'a => 'b. cont f}"
29063
7619f0561cd7 pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents: 29049
diff changeset
    24
by (simp_all add: Ex_cont adm_cont)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    25
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    26
syntax (xsymbols)
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    27
  "->"     :: "[type, type] => type"      ("(_ \<rightarrow>/ _)" [1,0]0)
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    28
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    29
notation
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    30
  Rep_CFun  ("(_$/_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    31
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    32
notation (xsymbols)
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    33
  Rep_CFun  ("(_\<cdot>/_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    34
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    35
notation (HTML output)
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    36
  Rep_CFun  ("(_\<cdot>/_)" [999,1000] 999)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    37
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    38
subsection {* Syntax for continuous lambda abstraction *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    39
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    40
syntax "_cabs" :: "'a"
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    41
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    42
parse_translation {*
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    43
(* rewrites (_cabs x t) => (Abs_CFun (%x. t)) *)
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    44
  [mk_binder_tr ("_cabs", @{const_syntax Abs_CFun})];
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    45
*}
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    46
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    47
text {* To avoid eta-contraction of body: *}
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    48
typed_print_translation {*
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    49
  let
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    50
    fun cabs_tr' _ _ [Abs abs] = let
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    51
          val (x,t) = atomic_abs_tr' abs
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    52
        in Syntax.const "_cabs" $ x $ t end
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    53
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    54
      | cabs_tr' _ T [t] = let
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    55
          val xT = domain_type (domain_type T);
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    56
          val abs' = ("x",xT,(incr_boundvars 1 t)$Bound 0);
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    57
          val (x,t') = atomic_abs_tr' abs';
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    58
        in Syntax.const "_cabs" $ x $ t' end;
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    59
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    60
  in [(@{const_syntax Abs_CFun}, cabs_tr')] end;
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    61
*}
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    62
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    63
text {* Syntax for nested abstractions *}
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    64
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    65
syntax
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    66
  "_Lambda" :: "[cargs, 'a] \<Rightarrow> logic"  ("(3LAM _./ _)" [1000, 10] 10)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    67
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    68
syntax (xsymbols)
25927
9c544dac6269 add space to binder syntax
huffman
parents: 25921
diff changeset
    69
  "_Lambda" :: "[cargs, 'a] \<Rightarrow> logic" ("(3\<Lambda> _./ _)" [1000, 10] 10)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    70
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    71
parse_ast_translation {*
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    72
(* rewrites (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    73
(* cf. Syntax.lambda_ast_tr from Syntax/syn_trans.ML *)
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    74
  let
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    75
    fun Lambda_ast_tr [pats, body] =
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    76
          Syntax.fold_ast_p "_cabs" (Syntax.unfold_ast "_cargs" pats, body)
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    77
      | Lambda_ast_tr asts = raise Syntax.AST ("Lambda_ast_tr", asts);
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    78
  in [("_Lambda", Lambda_ast_tr)] end;
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    79
*}
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    80
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    81
print_ast_translation {*
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    82
(* rewrites (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)
18078
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    83
(* cf. Syntax.abs_ast_tr' from Syntax/syn_trans.ML *)
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    84
  let
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    85
    fun cabs_ast_tr' asts =
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    86
      (case Syntax.unfold_ast_p "_cabs"
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    87
          (Syntax.Appl (Syntax.Constant "_cabs" :: asts)) of
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    88
        ([], _) => raise Syntax.AST ("cabs_ast_tr'", asts)
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    89
      | (xs, body) => Syntax.Appl
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    90
          [Syntax.Constant "_Lambda", Syntax.fold_ast "_cargs" xs, body]);
20e5a6440790 change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
huffman
parents: 18076
diff changeset
    91
  in [("_cabs", cabs_ast_tr')] end;
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    92
*}
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    93
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    94
text {* Dummy patterns for continuous abstraction *}
18079
9d4d70b175fd add translation for wildcard pattern
huffman
parents: 18078
diff changeset
    95
translations
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    96
  "\<Lambda> _. t" => "CONST Abs_CFun (\<lambda> _. t)"
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    97
18079
9d4d70b175fd add translation for wildcard pattern
huffman
parents: 18078
diff changeset
    98
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    99
subsection {* Continuous function space is pointed *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   100
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   101
lemma UU_CFun: "\<bottom> \<in> CFun"
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   102
by (simp add: CFun_def inst_fun_pcpo cont_const)
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   103
25827
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25786
diff changeset
   104
instance "->" :: (finite_po, finite_po) finite_po
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25786
diff changeset
   105
by (rule typedef_finite_po [OF type_definition_CFun])
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25786
diff changeset
   106
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25786
diff changeset
   107
instance "->" :: (finite_po, chfin) chfin
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   108
by (rule typedef_chfin [OF type_definition_CFun below_CFun_def])
25827
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25786
diff changeset
   109
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25927
diff changeset
   110
instance "->" :: (cpo, discrete_cpo) discrete_cpo
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   111
by intro_classes (simp add: below_CFun_def Rep_CFun_inject)
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25927
diff changeset
   112
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   113
instance "->" :: (cpo, pcpo) pcpo
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   114
by (rule typedef_pcpo [OF type_definition_CFun below_CFun_def UU_CFun])
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   115
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   116
lemmas Rep_CFun_strict =
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   117
  typedef_Rep_strict [OF type_definition_CFun below_CFun_def UU_CFun]
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   118
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   119
lemmas Abs_CFun_strict =
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   120
  typedef_Abs_strict [OF type_definition_CFun below_CFun_def UU_CFun]
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   121
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   122
text {* function application is strict in its first argument *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   123
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   124
lemma Rep_CFun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   125
by (simp add: Rep_CFun_strict)
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   126
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   127
text {* for compatibility with old HOLCF-Version *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   128
lemma inst_cfun_pcpo: "\<bottom> = (\<Lambda> x. \<bottom>)"
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   129
by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict)
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   130
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   131
subsection {* Basic properties of continuous functions *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   132
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   133
text {* Beta-equality for continuous functions *}
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   134
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   135
lemma Abs_CFun_inverse2: "cont f \<Longrightarrow> Rep_CFun (Abs_CFun f) = f"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   136
by (simp add: Abs_CFun_inverse CFun_def)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   137
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   138
lemma beta_cfun [simp]: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   139
by (simp add: Abs_CFun_inverse2)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   140
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   141
text {* Eta-equality for continuous functions *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   142
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   143
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   144
by (rule Rep_CFun_inverse)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   145
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   146
text {* Extensionality for continuous functions *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   147
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   148
lemma expand_cfun_eq: "(f = g) = (\<forall>x. f\<cdot>x = g\<cdot>x)"
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   149
by (simp add: Rep_CFun_inject [symmetric] expand_fun_eq)
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   150
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   151
lemma ext_cfun: "(\<And>x. f\<cdot>x = g\<cdot>x) \<Longrightarrow> f = g"
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   152
by (simp add: expand_cfun_eq)
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   153
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   154
text {* Extensionality wrt. ordering for continuous functions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   155
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   156
lemma expand_cfun_below: "f \<sqsubseteq> g = (\<forall>x. f\<cdot>x \<sqsubseteq> g\<cdot>x)" 
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   157
by (simp add: below_CFun_def expand_fun_below)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   158
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   159
lemma below_cfun_ext: "(\<And>x. f\<cdot>x \<sqsubseteq> g\<cdot>x) \<Longrightarrow> f \<sqsubseteq> g"
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   160
by (simp add: expand_cfun_below)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   161
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   162
text {* Congruence for continuous function application *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   163
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   164
lemma cfun_cong: "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> f\<cdot>x = g\<cdot>y"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   165
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   166
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   167
lemma cfun_fun_cong: "f = g \<Longrightarrow> f\<cdot>x = g\<cdot>x"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   168
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   169
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   170
lemma cfun_arg_cong: "x = y \<Longrightarrow> f\<cdot>x = f\<cdot>y"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   171
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   172
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   173
subsection {* Continuity of application *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   174
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   175
lemma cont_Rep_CFun1: "cont (\<lambda>f. f\<cdot>x)"
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   176
by (rule cont_Rep_CFun [THEN cont2cont_fun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   177
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   178
lemma cont_Rep_CFun2: "cont (\<lambda>x. f\<cdot>x)"
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   179
apply (cut_tac x=f in Rep_CFun)
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   180
apply (simp add: CFun_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   181
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   182
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   183
lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   184
lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   185
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   186
lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   187
lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   188
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   189
lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   190
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   191
text {* contlub, cont properties of @{term Rep_CFun} in each argument *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   192
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   193
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   194
by (rule contlub_Rep_CFun2 [THEN contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   195
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   196
lemma cont_cfun_arg: "chain Y \<Longrightarrow> range (\<lambda>i. f\<cdot>(Y i)) <<| f\<cdot>(\<Squnion>i. Y i)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   197
by (rule cont_Rep_CFun2 [THEN contE])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   198
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   199
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   200
by (rule contlub_Rep_CFun1 [THEN contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   201
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   202
lemma cont_cfun_fun: "chain F \<Longrightarrow> range (\<lambda>i. F i\<cdot>x) <<| (\<Squnion>i. F i)\<cdot>x"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   203
by (rule cont_Rep_CFun1 [THEN contE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   204
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   205
text {* monotonicity of application *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   206
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   207
lemma monofun_cfun_fun: "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   208
by (simp add: expand_cfun_below)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   209
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   210
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   211
by (rule monofun_Rep_CFun2 [THEN monofunE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   212
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   213
lemma monofun_cfun: "\<lbrakk>f \<sqsubseteq> g; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>y"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   214
by (rule below_trans [OF monofun_cfun_fun monofun_cfun_arg])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   215
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   216
text {* ch2ch - rules for the type @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   217
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   218
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   219
by (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   220
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   221
lemma ch2ch_Rep_CFunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   222
by (rule monofun_Rep_CFun2 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   223
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   224
lemma ch2ch_Rep_CFunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   225
by (rule monofun_Rep_CFun1 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   226
18076
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   227
lemma ch2ch_Rep_CFun [simp]:
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   228
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>(Y i))"
25884
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   229
by (simp add: chain_def monofun_cfun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   230
25884
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   231
lemma ch2ch_LAM [simp]:
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   232
  "\<lbrakk>\<And>x. chain (\<lambda>i. S i x); \<And>i. cont (\<lambda>x. S i x)\<rbrakk> \<Longrightarrow> chain (\<lambda>i. \<Lambda> x. S i x)"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   233
by (simp add: chain_def expand_cfun_below)
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   234
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   235
text {* contlub, cont properties of @{term Rep_CFun} in both arguments *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   236
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   237
lemma contlub_cfun: 
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   238
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. F i\<cdot>(Y i))"
18076
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   239
by (simp add: contlub_cfun_fun contlub_cfun_arg diag_lub)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   240
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   241
lemma cont_cfun: 
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   242
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> range (\<lambda>i. F i\<cdot>(Y i)) <<| (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   243
apply (rule thelubE)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   244
apply (simp only: ch2ch_Rep_CFun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   245
apply (simp only: contlub_cfun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   246
done
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   247
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   248
lemma contlub_LAM:
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   249
  "\<lbrakk>\<And>x. chain (\<lambda>i. F i x); \<And>i. cont (\<lambda>x. F i x)\<rbrakk>
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   250
    \<Longrightarrow> (\<Lambda> x. \<Squnion>i. F i x) = (\<Squnion>i. \<Lambda> x. F i x)"
25884
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   251
apply (simp add: thelub_CFun)
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   252
apply (simp add: Abs_CFun_inverse2)
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   253
apply (simp add: thelub_fun ch2ch_lambda)
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   254
done
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   255
25901
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   256
lemmas lub_distribs = 
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   257
  contlub_cfun [symmetric]
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   258
  contlub_LAM [symmetric]
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   259
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   260
text {* strictness *}
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   261
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   262
lemma strictI: "f\<cdot>x = \<bottom> \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   263
apply (rule UU_I)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   264
apply (erule subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   265
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   266
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   267
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   268
text {* the lub of a chain of continous functions is monotone *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   269
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   270
lemma lub_cfun_mono: "chain F \<Longrightarrow> monofun (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   271
apply (drule ch2ch_monofun [OF monofun_Rep_CFun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   272
apply (simp add: thelub_fun [symmetric])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   273
apply (erule monofun_lub_fun)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   274
apply (simp add: monofun_Rep_CFun2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   275
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   276
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents: 16314
diff changeset
   277
text {* a lemma about the exchange of lubs for type @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   278
16699
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
   279
lemma ex_lub_cfun:
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
   280
  "\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>j. \<Squnion>i. F j\<cdot>(Y i)) = (\<Squnion>i. \<Squnion>j. F j\<cdot>(Y i))"
18076
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   281
by (simp add: diag_lub)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   282
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   283
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
   284
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   285
lemma cont_lub_cfun: "chain F \<Longrightarrow> cont (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   286
apply (rule cont2cont_lub)
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   287
apply (erule monofun_Rep_CFun [THEN ch2ch_monofun])
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   288
apply (rule cont_Rep_CFun2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   289
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   290
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   291
text {* type @{typ "'a -> 'b"} is chain complete *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   292
16920
ded12c9e88c2 cleaned up
huffman
parents: 16699
diff changeset
   293
lemma lub_cfun: "chain F \<Longrightarrow> range F <<| (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
ded12c9e88c2 cleaned up
huffman
parents: 16699
diff changeset
   294
by (simp only: contlub_cfun_fun [symmetric] eta_cfun thelubE)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   295
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   296
lemma thelub_cfun: "chain F \<Longrightarrow> (\<Squnion>i. F i) = (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
16920
ded12c9e88c2 cleaned up
huffman
parents: 16699
diff changeset
   297
by (rule lub_cfun [THEN thelubI])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   298
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   299
subsection {* Continuity simplification procedure *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   300
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   301
text {* cont2cont lemma for @{term Rep_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   302
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   303
lemma cont2cont_Rep_CFun [cont2cont]:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   304
  assumes f: "cont (\<lambda>x. f x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   305
  assumes t: "cont (\<lambda>x. t x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   306
  shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   307
proof -
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   308
  have "cont (\<lambda>x. Rep_CFun (f x))"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   309
    using cont_Rep_CFun f by (rule cont2cont_app3)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   310
  thus "cont (\<lambda>x. (f x)\<cdot>(t x))"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   311
    using cont_Rep_CFun2 t by (rule cont2cont_app2)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   312
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   313
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   314
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
   315
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   316
lemma cont2mono_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   317
  "\<lbrakk>\<And>x. cont (\<lambda>y. f x y); \<And>y. monofun (\<lambda>x. f x y)\<rbrakk>
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   318
    \<Longrightarrow> monofun (\<lambda>x. \<Lambda> y. f x y)"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   319
  unfolding monofun_def expand_cfun_below by simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   320
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   321
text {* cont2cont Lemma for @{term "%x. LAM y. f x y"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   322
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   323
text {*
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   324
  Not suitable as a cont2cont rule, because on nested lambdas
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   325
  it causes exponential blow-up in the number of subgoals.
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   326
*}
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   327
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   328
lemma cont2cont_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   329
  assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   330
  assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   331
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   332
proof (rule cont_Abs_CFun)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   333
  fix x
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   334
  from f1 show "f x \<in> CFun" by (simp add: CFun_def)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   335
  from f2 show "cont f" by (rule cont2cont_lambda)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   336
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   337
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   338
text {*
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   339
  This version does work as a cont2cont rule, since it
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   340
  has only a single subgoal.
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   341
*}
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   342
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   343
lemma cont2cont_LAM' [cont2cont]:
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   344
  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   345
  assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   346
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   347
proof (rule cont2cont_LAM)
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29533
diff changeset
   348
  fix x :: 'a show "cont (\<lambda>y. f x y)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29533
diff changeset
   349
    using f by (rule cont_fst_snd_D2)
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   350
next
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29533
diff changeset
   351
  fix y :: 'b show "cont (\<lambda>x. f x y)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 29533
diff changeset
   352
    using f by (rule cont_fst_snd_D1)
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   353
qed
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   354
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   355
lemma cont2cont_LAM_discrete [cont2cont]:
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   356
  "(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   357
by (simp add: cont2cont_LAM)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   358
16055
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   359
lemmas cont_lemmas1 =
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   360
  cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
58186c507750 moved continuity simproc to Cfun.thy
huffman
parents: 15641
diff changeset
   361
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   362
subsection {* Miscellaneous *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   363
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   364
text {* Monotonicity of @{term Abs_CFun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   365
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   366
lemma semi_monofun_Abs_CFun:
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   367
  "\<lbrakk>cont f; cont g; f \<sqsubseteq> g\<rbrakk> \<Longrightarrow> Abs_CFun f \<sqsubseteq> Abs_CFun g"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   368
by (simp add: below_CFun_def Abs_CFun_inverse2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   369
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   370
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
   371
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   372
lemma chfin_Rep_CFunR: "chain (Y::nat => 'a::cpo->'b::chfin)  
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   373
      ==> !s. ? n. (LUB i. Y i)$s = Y n$s"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   374
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   375
apply (subst contlub_cfun_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   376
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   377
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
   378
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   379
18089
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   380
lemma adm_chfindom: "adm (\<lambda>(u::'a::cpo \<rightarrow> 'b::chfin). P(u\<cdot>s))"
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   381
by (rule adm_subst, simp, rule adm_chfin)
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   382
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   383
subsection {* Continuous injection-retraction pairs *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   384
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   385
text {* Continuous retractions are strict. *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   386
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   387
lemma retraction_strict:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   388
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   389
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   390
apply (drule_tac x="\<bottom>" in spec)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   391
apply (erule subst)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   392
apply (rule monofun_cfun_arg)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   393
apply (rule minimal)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   394
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   395
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   396
lemma injection_eq:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   397
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x = g\<cdot>y) = (x = y)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   398
apply (rule iffI)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   399
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   400
apply simp
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   401
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   402
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   403
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 31041
diff changeset
   404
lemma injection_below:
16314
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   405
  "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x \<sqsubseteq> g\<cdot>y) = (x \<sqsubseteq> y)"
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   406
apply (rule iffI)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   407
apply (drule_tac f=f in monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   408
apply simp
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   409
apply (erule monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   410
done
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   411
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   412
lemma injection_defined_rev:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   413
  "\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; g\<cdot>z = \<bottom>\<rbrakk> \<Longrightarrow> z = \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   414
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   415
apply (simp add: retraction_strict)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   416
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   417
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   418
lemma injection_defined:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   419
  "\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; z \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> g\<cdot>z \<noteq> \<bottom>"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   420
by (erule contrapos_nn, rule injection_defined_rev)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   421
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   422
text {* propagation of flatness and chain-finiteness by retractions *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   423
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   424
lemma chfin2chfin:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   425
  "\<forall>y. (f::'a::chfin \<rightarrow> 'b)\<cdot>(g\<cdot>y) = y
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   426
    \<Longrightarrow> \<forall>Y::nat \<Rightarrow> 'b. chain Y \<longrightarrow> (\<exists>n. max_in_chain n Y)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   427
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   428
apply (drule_tac f=g in chain_monofun)
25921
0ca392ab7f37 change class axiom chfin to rule_format
huffman
parents: 25920
diff changeset
   429
apply (drule chfin)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   430
apply (unfold max_in_chain_def)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   431
apply (simp add: injection_eq)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   432
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   433
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   434
lemma flat2flat:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   435
  "\<forall>y. (f::'a::flat \<rightarrow> 'b::pcpo)\<cdot>(g\<cdot>y) = y
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   436
    \<Longrightarrow> \<forall>x y::'b. x \<sqsubseteq> y \<longrightarrow> x = \<bottom> \<or> x = y"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   437
apply clarify
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   438
apply (drule_tac f=g in monofun_cfun_arg)
25920
8df5eabda5f6 change class axiom ax_flat to rule_format
huffman
parents: 25901
diff changeset
   439
apply (drule ax_flat)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   440
apply (erule disjE)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   441
apply (simp add: injection_defined_rev)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   442
apply (simp add: injection_eq)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   443
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   444
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   445
text {* a result about functions with flat codomain *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   446
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   447
lemma flat_eqI: "\<lbrakk>(x::'a::flat) \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> x = y"
25920
8df5eabda5f6 change class axiom ax_flat to rule_format
huffman
parents: 25901
diff changeset
   448
by (drule ax_flat, simp)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   449
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   450
lemma flat_codom:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   451
  "f\<cdot>x = (c::'b::flat) \<Longrightarrow> f\<cdot>\<bottom> = \<bottom> \<or> (\<forall>z. f\<cdot>z = c)"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   452
apply (case_tac "f\<cdot>x = \<bottom>")
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   453
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   454
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   455
apply (erule_tac t="\<bottom>" in subst)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   456
apply (rule minimal [THEN monofun_cfun_arg])
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   457
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   458
apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   459
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   460
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   461
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   462
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   463
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   464
subsection {* Identity and composition *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   465
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   466
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   467
  ID :: "'a \<rightarrow> 'a" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   468
  "ID = (\<Lambda> x. x)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   469
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   470
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   471
  cfcomp  :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   472
  oo_def: "cfcomp = (\<Lambda> f g x. f\<cdot>(g\<cdot>x))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   473
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   474
abbreviation
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   475
  cfcomp_syn :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c"  (infixr "oo" 100)  where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   476
  "f oo g == cfcomp\<cdot>f\<cdot>g"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   477
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   478
lemma ID1 [simp]: "ID\<cdot>x = x"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   479
by (simp add: ID_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   480
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   481
lemma cfcomp1: "(f oo g) = (\<Lambda> x. f\<cdot>(g\<cdot>x))"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   482
by (simp add: oo_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   483
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   484
lemma cfcomp2 [simp]: "(f oo g)\<cdot>x = f\<cdot>(g\<cdot>x)"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   485
by (simp add: cfcomp1)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   486
27274
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   487
lemma cfcomp_LAM: "cont g \<Longrightarrow> f oo (\<Lambda> x. g x) = (\<Lambda> x. f\<cdot>(g x))"
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   488
by (simp add: cfcomp1)
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   489
19709
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   490
lemma cfcomp_strict [simp]: "\<bottom> oo f = \<bottom>"
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   491
by (simp add: expand_cfun_eq)
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   492
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   493
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   494
  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
   495
  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
   496
  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
   497
  The identity arrow is interpretation of @{term ID}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   498
  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
   499
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   500
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   501
lemma ID2 [simp]: "f oo ID = f"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   502
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   503
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   504
lemma ID3 [simp]: "ID oo f = f"
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   505
by (rule ext_cfun, simp)
15576
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 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
   508
by (rule ext_cfun, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   509
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   510
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   511
subsection {* Strictified functions *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   512
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   513
defaultsort pcpo
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   514
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   515
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   516
  strictify  :: "('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'b" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   517
  "strictify = (\<Lambda> f x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   518
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   519
text {* results about strictify *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   520
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   521
lemma cont_strictify1: "cont (\<lambda>f. if x = \<bottom> then \<bottom> else f\<cdot>x)"
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   522
by (simp add: cont_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   523
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   524
lemma monofun_strictify2: "monofun (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   525
apply (rule monofunI)
25786
6b3c79acac1f move lemmas from Cont.thy to Ffun.thy;
huffman
parents: 25723
diff changeset
   526
apply (auto simp add: monofun_cfun_arg)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   527
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   528
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   529
(*FIXME: long proof*)
25723
80c06e4d4db6 move bottom-related stuff back into Pcpo.thy
huffman
parents: 25701
diff changeset
   530
lemma contlub_strictify2: "contlub (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   531
apply (rule contlubI)
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   532
apply (case_tac "(\<Squnion>i. Y i) = \<bottom>")
16699
24b494ff8f0f use new pcpodef package
huffman
parents: 16417
diff changeset
   533
apply (drule (1) chain_UU_I)
18076
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   534
apply simp
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   535
apply (simp del: if_image_distrib)
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   536
apply (simp only: contlub_cfun_arg)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   537
apply (rule lub_equal2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   538
apply (rule chain_mono2 [THEN exE])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   539
apply (erule chain_UU_I_inverse2)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   540
apply (assumption)
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   541
apply (rule_tac x=x in exI, clarsimp)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   542
apply (erule chain_monofun)
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   543
apply (erule monofun_strictify2 [THEN ch2ch_monofun])
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   544
done
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   545
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   546
lemmas cont_strictify2 =
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   547
  monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   548
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   549
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   550
  unfolding strictify_def
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   551
  by (simp add: cont_strictify1 cont_strictify2 cont2cont_LAM)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   552
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   553
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   554
by (simp add: strictify_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   555
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   556
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   557
by (simp add: strictify_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   558
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   559
subsection {* Continuous let-bindings *}
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   560
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   561
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   562
  CLet :: "'a \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'b" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   563
  "CLet = (\<Lambda> s f. f\<cdot>s)"
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   564
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   565
syntax
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   566
  "_CLet" :: "[letbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10)
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   567
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   568
translations
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   569
  "_CLet (_binds b bs) e" == "_CLet b (_CLet bs e)"
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   570
  "Let x = a in e" == "CONST CLet\<cdot>a\<cdot>(\<Lambda> x. e)"
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   571
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   572
end