src/HOLCF/Cfun.thy
author huffman
Sat, 27 Nov 2010 12:55:12 -0800
changeset 40770 6023808b38d4
parent 40767 a3e505b236e7
child 40771 1c6f7d4b110e
permissions -rw-r--r--
rename cpodef theorems: lub_foo -> is_lub_foo, thelub_foo -> lub_foo
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
35794
8cd7134275cc use headers consistently
huffman
parents: 35641
diff changeset
     3
    Author:     Brian Huffman
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     4
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     5
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     6
header {* The type of continuous functions *}
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     7
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
theory Cfun
40001
666c3751227c rename Ffun.thy to Fun_Cpo.thy
huffman
parents: 39985
diff changeset
     9
imports Pcpodef Fun_Cpo Product_Cpo
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    11
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35933
diff changeset
    12
default_sort cpo
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    13
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    14
subsection {* Definition of continuous function type *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    15
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    16
cpodef ('a, 'b) cfun (infixr "->" 0) = "{f::'a => 'b. cont f}"
40011
b974cf829099 cleaned up Fun_Cpo.thy; deprecated a few theorem names
huffman
parents: 40008
diff changeset
    17
by (auto intro: cont_const adm_cont)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    18
35427
ad039d29e01c proper (type_)notation;
wenzelm
parents: 35168
diff changeset
    19
type_notation (xsymbols)
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 35168
diff changeset
    20
  cfun  ("(_ \<rightarrow>/ _)" [1, 0] 0)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    21
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    22
notation
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    23
  Rep_cfun  ("(_$/_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    24
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    25
notation (xsymbols)
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    26
  Rep_cfun  ("(_\<cdot>/_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    27
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    28
notation (HTML output)
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    29
  Rep_cfun  ("(_\<cdot>/_)" [999,1000] 999)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    30
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    31
subsection {* Syntax for continuous lambda abstraction *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    32
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
    33
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
    34
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
    35
parse_translation {*
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    36
(* rewrite (_cabs x t) => (Abs_cfun (%x. t)) *)
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    37
  [mk_binder_tr (@{syntax_const "_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
    38
*}
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    39
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    40
text {* To avoid eta-contraction of body: *}
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    41
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
    42
  let
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    43
    fun cabs_tr' _ _ [Abs abs] = let
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    44
          val (x,t) = atomic_abs_tr' abs
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    45
        in Syntax.const @{syntax_const "_cabs"} $ x $ t end
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    46
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    47
      | cabs_tr' _ T [t] = let
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    48
          val xT = domain_type (domain_type T);
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    49
          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
    50
          val (x,t') = atomic_abs_tr' abs';
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    51
        in Syntax.const @{syntax_const "_cabs"} $ x $ t' end;
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    52
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    53
  in [(@{const_syntax Abs_cfun}, cabs_tr')] end;
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    54
*}
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    55
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    56
text {* Syntax for nested abstractions *}
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    57
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    58
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
    59
  "_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
    60
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    61
syntax (xsymbols)
25927
9c544dac6269 add space to binder syntax
huffman
parents: 25921
diff changeset
    62
  "_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
    63
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    64
parse_ast_translation {*
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    65
(* rewrite (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    66
(* cf. Syntax.lambda_ast_tr from src/Pure/Syntax/syn_trans.ML *)
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
    67
  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
    68
    fun Lambda_ast_tr [pats, body] =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    69
          Syntax.fold_ast_p @{syntax_const "_cabs"}
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    70
            (Syntax.unfold_ast @{syntax_const "_cargs"} pats, body)
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
    71
      | Lambda_ast_tr asts = raise Syntax.AST ("Lambda_ast_tr", asts);
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    72
  in [(@{syntax_const "_Lambda"}, Lambda_ast_tr)] end;
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    73
*}
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    74
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    75
print_ast_translation {*
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    76
(* rewrite (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    77
(* cf. Syntax.abs_ast_tr' from src/Pure/Syntax/syn_trans.ML *)
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
    78
  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
    79
    fun cabs_ast_tr' asts =
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    80
      (case Syntax.unfold_ast_p @{syntax_const "_cabs"}
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    81
          (Syntax.Appl (Syntax.Constant @{syntax_const "_cabs"} :: asts)) of
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
    82
        ([], _) => 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
    83
      | (xs, body) => Syntax.Appl
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    84
          [Syntax.Constant @{syntax_const "_Lambda"},
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    85
           Syntax.fold_ast @{syntax_const "_cargs"} xs, body]);
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    86
  in [(@{syntax_const "_cabs"}, cabs_ast_tr')] end
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    87
*}
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    88
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    89
text {* Dummy patterns for continuous abstraction *}
18079
9d4d70b175fd add translation for wildcard pattern
huffman
parents: 18078
diff changeset
    90
translations
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    91
  "\<Lambda> _. t" => "CONST Abs_cfun (\<lambda> _. t)"
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    92
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    93
subsection {* Continuous function space is pointed *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    94
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    95
lemma UU_cfun: "\<bottom> \<in> cfun"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    96
by (simp add: cfun_def inst_fun_pcpo)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    97
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 35168
diff changeset
    98
instance cfun :: (cpo, discrete_cpo) discrete_cpo
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    99
by intro_classes (simp add: below_cfun_def Rep_cfun_inject)
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25927
diff changeset
   100
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 35168
diff changeset
   101
instance cfun :: (cpo, pcpo) pcpo
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   102
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
   103
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   104
lemmas Rep_cfun_strict =
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   105
  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
   106
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   107
lemmas Abs_cfun_strict =
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   108
  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
   109
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   110
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
   111
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   112
lemma Rep_cfun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   113
by (simp add: Rep_cfun_strict)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   114
35641
a17bc4cec23a add simp rule LAM_strict
huffman
parents: 35547
diff changeset
   115
lemma LAM_strict [simp]: "(\<Lambda> x. \<bottom>) = \<bottom>"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   116
by (simp add: inst_fun_pcpo [symmetric] Abs_cfun_strict)
35641
a17bc4cec23a add simp rule LAM_strict
huffman
parents: 35547
diff changeset
   117
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   118
text {* for compatibility with old HOLCF-Version *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   119
lemma inst_cfun_pcpo: "\<bottom> = (\<Lambda> x. \<bottom>)"
35641
a17bc4cec23a add simp rule LAM_strict
huffman
parents: 35547
diff changeset
   120
by simp
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   121
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   122
subsection {* Basic properties of continuous functions *}
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
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
   125
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   126
lemma Abs_cfun_inverse2: "cont f \<Longrightarrow> Rep_cfun (Abs_cfun f) = f"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   127
by (simp add: Abs_cfun_inverse cfun_def)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   128
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   129
lemma beta_cfun: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   130
by (simp add: Abs_cfun_inverse2)
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
   131
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   132
text {* Beta-reduction simproc *}
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   133
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   134
text {*
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   135
  Given the term @{term "(\<Lambda> x. f x)\<cdot>y"}, the procedure tries to
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   136
  construct the theorem @{term "(\<Lambda> x. f x)\<cdot>y == f y"}.  If this
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   137
  theorem cannot be completely solved by the cont2cont rules, then
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   138
  the procedure returns the ordinary conditional @{text beta_cfun}
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   139
  rule.
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   140
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   141
  The simproc does not solve any more goals that would be solved by
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   142
  using @{text beta_cfun} as a simp rule.  The advantage of the
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   143
  simproc is that it can avoid deeply-nested calls to the simplifier
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   144
  that would otherwise be caused by large continuity side conditions.
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   145
*}
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   146
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   147
simproc_setup beta_cfun_proc ("Abs_cfun f\<cdot>x") = {*
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   148
  fn phi => fn ss => fn ct =>
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   149
    let
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   150
      val dest = Thm.dest_comb;
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   151
      val (f, x) = (apfst (snd o dest o snd o dest) o dest) ct;
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   152
      val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   153
      val tr = instantiate' [SOME T, SOME U] [SOME f, SOME x]
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   154
          (mk_meta_eq @{thm beta_cfun});
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   155
      val rules = Cont2ContData.get (Simplifier.the_context ss);
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   156
      val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   157
    in SOME (perhaps (SINGLE (tac 1)) tr) end
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   158
*}
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   159
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
   160
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
   161
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   162
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   163
by (rule Rep_cfun_inverse)
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
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   165
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
   166
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   167
lemma cfun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f\<cdot>x = g\<cdot>x)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   168
by (simp add: Rep_cfun_inject [symmetric] fun_eq_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   169
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   170
lemma cfun_eqI: "(\<And>x. f\<cdot>x = g\<cdot>x) \<Longrightarrow> f = g"
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   171
by (simp add: cfun_eq_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   172
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   173
text {* Extensionality wrt. ordering for continuous functions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   174
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   175
lemma cfun_below_iff: "f \<sqsubseteq> g \<longleftrightarrow> (\<forall>x. f\<cdot>x \<sqsubseteq> g\<cdot>x)" 
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   176
by (simp add: below_cfun_def fun_below_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   177
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   178
lemma cfun_belowI: "(\<And>x. f\<cdot>x \<sqsubseteq> g\<cdot>x) \<Longrightarrow> f \<sqsubseteq> g"
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   179
by (simp add: cfun_below_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   180
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   181
text {* Congruence for continuous function application *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
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
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
   184
by simp
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
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
   187
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   188
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
   189
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
   190
by simp
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   191
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
   192
subsection {* Continuity of application *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   193
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   194
lemma cont_Rep_cfun1: "cont (\<lambda>f. f\<cdot>x)"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   195
by (rule cont_Rep_cfun [THEN cont2cont_fun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   196
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   197
lemma cont_Rep_cfun2: "cont (\<lambda>x. f\<cdot>x)"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   198
apply (cut_tac x=f in Rep_cfun)
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   199
apply (simp add: cfun_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   200
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   201
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   202
lemmas monofun_Rep_cfun = cont_Rep_cfun [THEN cont2mono]
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   203
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   204
lemmas monofun_Rep_cfun1 = cont_Rep_cfun1 [THEN cont2mono, standard]
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   205
lemmas monofun_Rep_cfun2 = cont_Rep_cfun2 [THEN cont2mono, standard]
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   206
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   207
text {* contlub, cont properties of @{term Rep_cfun} in each argument *}
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
   208
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   209
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   210
by (rule cont_Rep_cfun2 [THEN cont2contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   211
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   212
lemma cont_cfun_arg: "chain Y \<Longrightarrow> range (\<lambda>i. f\<cdot>(Y i)) <<| f\<cdot>(\<Squnion>i. Y i)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   213
by (rule cont_Rep_cfun2 [THEN contE])
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
   214
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   215
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   216
by (rule cont_Rep_cfun1 [THEN cont2contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   217
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   218
lemma cont_cfun_fun: "chain F \<Longrightarrow> range (\<lambda>i. F i\<cdot>x) <<| (\<Squnion>i. F i)\<cdot>x"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   219
by (rule cont_Rep_cfun1 [THEN contE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   220
16209
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   221
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
   222
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   223
lemma monofun_cfun_fun: "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>x"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   224
by (simp add: cfun_below_iff)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   225
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
   226
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   227
by (rule monofun_Rep_cfun2 [THEN monofunE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   228
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
   229
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
   230
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
   231
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
   232
text {* ch2ch - rules for the type @{typ "'a -> 'b"} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   233
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
   234
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   235
by (erule monofun_Rep_cfun2 [THEN ch2ch_monofun])
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
   236
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   237
lemma ch2ch_Rep_cfunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   238
by (rule monofun_Rep_cfun2 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   239
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   240
lemma ch2ch_Rep_cfunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   241
by (rule monofun_Rep_cfun1 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   242
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   243
lemma ch2ch_Rep_cfun [simp]:
18076
e2e626b673b3 cleaned up; ch2ch_Rep_CFun is a simp rule
huffman
parents: 17832
diff changeset
   244
  "\<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
   245
by (simp add: chain_def monofun_cfun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   246
25884
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   247
lemma ch2ch_LAM [simp]:
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   248
  "\<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)"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   249
by (simp add: chain_def cfun_below_iff)
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   250
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   251
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
   252
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
   253
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
   254
  "\<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
   255
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
   256
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
   257
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
   258
  "\<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
   259
apply (rule thelubE)
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   260
apply (simp only: ch2ch_Rep_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
   261
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
   262
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
   263
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   264
lemma contlub_LAM:
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   265
  "\<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
   266
    \<Longrightarrow> (\<Lambda> x. \<Squnion>i. F i x) = (\<Squnion>i. \<Lambda> x. F i x)"
40770
6023808b38d4 rename cpodef theorems: lub_foo -> is_lub_foo, thelub_foo -> lub_foo
huffman
parents: 40767
diff changeset
   267
apply (simp add: lub_cfun)
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   268
apply (simp add: Abs_cfun_inverse2)
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   269
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
   270
done
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   271
25901
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   272
lemmas lub_distribs = 
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   273
  contlub_cfun [symmetric]
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   274
  contlub_LAM [symmetric]
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   275
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
   276
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
   277
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
   278
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
   279
apply (rule UU_I)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   280
apply (erule subst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   281
apply (rule minimal [THEN monofun_cfun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   282
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   283
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   284
text {* type @{typ "'a -> 'b"} is chain complete *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   285
16920
ded12c9e88c2 cleaned up
huffman
parents: 16699
diff changeset
   286
lemma lub_cfun: "chain F \<Longrightarrow> range F <<| (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
ded12c9e88c2 cleaned up
huffman
parents: 16699
diff changeset
   287
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
   288
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   289
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
   290
by (rule lub_cfun [THEN thelubI])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   291
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   292
subsection {* Continuity simplification procedure *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   293
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   294
text {* cont2cont lemma for @{term Rep_cfun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   295
40326
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   296
lemma cont2cont_APP [simp, cont2cont]:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   297
  assumes f: "cont (\<lambda>x. f x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   298
  assumes t: "cont (\<lambda>x. t x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   299
  shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   300
proof -
40006
116e94f9543b remove unneeded lemmas from Fun_Cpo.thy
huffman
parents: 40005
diff changeset
   301
  have 1: "\<And>y. cont (\<lambda>x. (f x)\<cdot>y)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   302
    using cont_Rep_cfun1 f by (rule cont_compose)
40006
116e94f9543b remove unneeded lemmas from Fun_Cpo.thy
huffman
parents: 40005
diff changeset
   303
  show "cont (\<lambda>x. (f x)\<cdot>(t x))"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   304
    using t cont_Rep_cfun2 1 by (rule cont_apply)
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   305
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   306
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   307
text {*
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   308
  Two specific lemmas for the combination of LCF and HOL terms.
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   309
  These lemmas are needed in theories that use types like @{typ "'a \<rightarrow> 'b \<Rightarrow> 'c"}.
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   310
*}
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   311
40326
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   312
lemma cont_APP_app [simp]: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. ((f x)\<cdot>(g x)) s)"
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   313
by (rule cont2cont_APP [THEN cont2cont_fun])
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   314
40326
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   315
lemma cont_APP_app_app [simp]: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (\<lambda>x. ((f x)\<cdot>(g x)) s t)"
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   316
by (rule cont_APP_app [THEN cont2cont_fun])
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   317
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   318
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   319
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
   320
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   321
lemma cont2mono_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   322
  "\<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
   323
    \<Longrightarrow> monofun (\<lambda>x. \<Lambda> y. f x y)"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   324
  unfolding monofun_def cfun_below_iff by simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   325
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   326
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
   327
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   328
text {*
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   329
  Not suitable as a cont2cont rule, because on nested lambdas
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   330
  it causes exponential blow-up in the number of subgoals.
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   331
*}
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   332
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   333
lemma cont2cont_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   334
  assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   335
  assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   336
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   337
proof (rule cont_Abs_cfun)
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   338
  fix x
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   339
  from f1 show "f x \<in> cfun" by (simp add: cfun_def)
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   340
  from f2 show "cont f" by (rule cont2cont_lambda)
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   341
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   342
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   343
text {*
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   344
  This version does work as a cont2cont rule, since it
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   345
  has only a single subgoal.
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   346
*}
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   347
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   348
lemma cont2cont_LAM' [simp, cont2cont]:
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   349
  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   350
  assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   351
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
39808
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39302
diff changeset
   352
using assms by (simp add: cont2cont_LAM prod_cont_iff)
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   353
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   354
lemma cont2cont_LAM_discrete [simp, cont2cont]:
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   355
  "(\<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
   356
by (simp add: cont2cont_LAM)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   357
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   358
subsection {* Miscellaneous *}
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   359
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   360
text {* Monotonicity of @{term Abs_cfun} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   361
40433
3128c2a54785 remove some unnecessary lemmas; move monofun_LAM to Cfun.thy
huffman
parents: 40327
diff changeset
   362
lemma monofun_LAM:
3128c2a54785 remove some unnecessary lemmas; move monofun_LAM to Cfun.thy
huffman
parents: 40327
diff changeset
   363
  "\<lbrakk>cont f; cont g; \<And>x. f x \<sqsubseteq> g x\<rbrakk> \<Longrightarrow> (\<Lambda> x. f x) \<sqsubseteq> (\<Lambda> x. g x)"
3128c2a54785 remove some unnecessary lemmas; move monofun_LAM to Cfun.thy
huffman
parents: 40327
diff changeset
   364
by (simp add: cfun_below_iff)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   365
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   366
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
   367
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   368
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
   369
      ==> !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
   370
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   371
apply (subst contlub_cfun_fun)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   372
apply assumption
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   373
apply (fast intro!: thelubI chfin lub_finch2 chfin2finch ch2ch_Rep_cfunL)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   374
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   375
18089
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   376
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
   377
by (rule adm_subst, simp, rule adm_chfin)
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   378
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   379
subsection {* Continuous injection-retraction pairs *}
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   380
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   381
text {* Continuous retractions are strict. *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   382
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   383
lemma retraction_strict:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   384
  "\<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
   385
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   386
apply (drule_tac x="\<bottom>" in spec)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   387
apply (erule subst)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   388
apply (rule monofun_cfun_arg)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   389
apply (rule minimal)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   390
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   391
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   392
lemma injection_eq:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   393
  "\<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
   394
apply (rule iffI)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   395
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   396
apply simp
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   397
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   398
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   399
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
   400
lemma injection_below:
16314
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   401
  "\<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
   402
apply (rule iffI)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   403
apply (drule_tac f=f in monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   404
apply simp
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   405
apply (erule monofun_cfun_arg)
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   406
done
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   407
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   408
lemma injection_defined_rev:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   409
  "\<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
   410
apply (drule_tac f=f in cfun_arg_cong)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   411
apply (simp add: retraction_strict)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   412
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   413
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   414
lemma injection_defined:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   415
  "\<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
   416
by (erule contrapos_nn, rule injection_defined_rev)
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   417
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   418
text {* a result about functions with flat codomain *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   419
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   420
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
   421
by (drule ax_flat, simp)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   422
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   423
lemma flat_codom:
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   424
  "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
   425
apply (case_tac "f\<cdot>x = \<bottom>")
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   426
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   427
apply (rule UU_I)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   428
apply (erule_tac t="\<bottom>" in subst)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   429
apply (rule minimal [THEN monofun_cfun_arg])
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   430
apply clarify
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   431
apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   432
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   433
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
   434
done
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   435
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   436
subsection {* Identity and composition *}
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   437
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   438
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   439
  ID :: "'a \<rightarrow> 'a" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   440
  "ID = (\<Lambda> x. x)"
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   441
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   442
definition
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   443
  cfcomp  :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c" where
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   444
  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
   445
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   446
abbreviation
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   447
  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
   448
  "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
   449
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   450
lemma ID1 [simp]: "ID\<cdot>x = x"
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   451
by (simp add: ID_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   452
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   453
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
   454
by (simp add: oo_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   455
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   456
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
   457
by (simp add: cfcomp1)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   458
27274
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   459
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
   460
by (simp add: cfcomp1)
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   461
19709
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   462
lemma cfcomp_strict [simp]: "\<bottom> oo f = \<bottom>"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   463
by (simp add: cfun_eq_iff)
19709
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   464
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   465
text {*
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   466
  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
   467
  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
   468
  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
   469
  The identity arrow is interpretation of @{term ID}.
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   470
  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
   471
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   472
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   473
lemma ID2 [simp]: "f oo ID = f"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   474
by (rule cfun_eqI, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   475
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   476
lemma ID3 [simp]: "ID oo f = f"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   477
by (rule cfun_eqI, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   478
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   479
lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 40001
diff changeset
   480
by (rule cfun_eqI, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   481
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   482
subsection {* Strictified functions *}
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   483
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35933
diff changeset
   484
default_sort pcpo
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   485
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
   486
definition
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   487
  seq :: "'a \<rightarrow> 'b \<rightarrow> 'b" where
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   488
  "seq = (\<Lambda> x. if x = \<bottom> then \<bottom> else ID)"
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   489
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   490
lemma cont_seq: "cont (\<lambda>x. if x = \<bottom> then \<bottom> else y)"
40046
ba2e41c8b725 introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman
parents: 40011
diff changeset
   491
unfolding cont_def is_lub_def is_ub_def ball_simps
ba2e41c8b725 introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman
parents: 40011
diff changeset
   492
by (simp add: lub_eq_bottom_iff)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   493
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   494
lemma seq_conv_if: "seq\<cdot>x = (if x = \<bottom> then \<bottom> else ID)"
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   495
unfolding seq_def by (simp add: cont_seq)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   496
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   497
lemma seq1 [simp]: "seq\<cdot>\<bottom> = \<bottom>"
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   498
by (simp add: seq_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   499
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   500
lemma seq2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x = ID"
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   501
by (simp add: seq_conv_if)
40046
ba2e41c8b725 introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman
parents: 40011
diff changeset
   502
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   503
lemma seq3 [simp]: "seq\<cdot>x\<cdot>\<bottom> = \<bottom>"
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   504
by (simp add: seq_conv_if)
40093
c2d36bc4cd14 add lemma strict3
huffman
parents: 40091
diff changeset
   505
c2d36bc4cd14 add lemma strict3
huffman
parents: 40091
diff changeset
   506
definition
40046
ba2e41c8b725 introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman
parents: 40011
diff changeset
   507
  strictify  :: "('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'b" where
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   508
  "strictify = (\<Lambda> f x. seq\<cdot>x\<cdot>(f\<cdot>x))"
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   509
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   510
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
40046
ba2e41c8b725 introduce function strict :: 'a -> 'b -> 'b, which works like Haskell's seq; use strict instead of strictify in various definitions
huffman
parents: 40011
diff changeset
   511
unfolding strictify_def by simp
16085
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
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   514
by (simp add: strictify_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   515
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   516
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
   517
by (simp add: strictify_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   518
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   519
subsection {* Continuity of let-bindings *}
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   520
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   521
lemma cont2cont_Let:
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   522
  assumes f: "cont (\<lambda>x. f x)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   523
  assumes g1: "\<And>y. cont (\<lambda>x. g x y)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   524
  assumes g2: "\<And>x. cont (\<lambda>y. g x y)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   525
  shows "cont (\<lambda>x. let y = f x in g x y)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   526
unfolding Let_def using f g2 g1 by (rule cont_apply)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   527
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   528
lemma cont2cont_Let' [simp, cont2cont]:
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   529
  assumes f: "cont (\<lambda>x. f x)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   530
  assumes g: "cont (\<lambda>p. g (fst p) (snd p))"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   531
  shows "cont (\<lambda>x. let y = f x in g x y)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   532
using f
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   533
proof (rule cont2cont_Let)
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   534
  fix x show "cont (\<lambda>y. g x y)"
40003
427106657e04 remove unused lemmas cont_fst_snd_D1, cont_fst_snd_D2
huffman
parents: 40002
diff changeset
   535
    using g by (simp add: prod_cont_iff)
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   536
next
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   537
  fix y show "cont (\<lambda>x. g x y)"
40003
427106657e04 remove unused lemmas cont_fst_snd_D1, cont_fst_snd_D2
huffman
parents: 40002
diff changeset
   538
    using g by (simp add: prod_cont_iff)
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   539
qed
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   540
39145
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   541
text {* The simple version (suggested by Joachim Breitner) is needed if
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   542
  the type of the defined term is not a cpo. *}
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   543
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   544
lemma cont2cont_Let_simple [simp, cont2cont]:
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   545
  assumes "\<And>y. cont (\<lambda>x. g x y)"
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   546
  shows "cont (\<lambda>x. let y = t in g x y)"
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   547
unfolding Let_def using assms .
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   548
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   549
end