src/HOL/HOLCF/Cfun.thy
author wenzelm
Sat, 17 Mar 2018 20:32:39 +0100
changeset 67895 cd00999d2d30
parent 67312 0d25e02759b7
child 68383 93a42bd62ede
permissions -rw-r--r--
more position information;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 42057
diff changeset
     1
(*  Title:      HOL/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
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
     6
section \<open>The type of continuous functions\<close>
15576
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
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
     9
  imports Cpodef 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
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    14
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    15
subsection \<open>Definition of continuous function type\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    16
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    17
definition "cfun = {f::'a \<Rightarrow> 'b. cont f}"
45695
b108b3d7c49e prefer cpodef without extra definition;
wenzelm
parents: 45606
diff changeset
    18
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    19
cpodef ('a, 'b) cfun ("(_ \<rightarrow>/ _)" [1, 0] 0) = "cfun :: ('a \<Rightarrow> 'b) set"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    20
  by (auto simp: cfun_def intro: cont_const adm_cont)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    21
61998
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    22
type_notation (ASCII)
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    23
  cfun  (infixr "->" 0)
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    24
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    25
notation (ASCII)
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    26
  Rep_cfun  ("(_$/_)" [999,1000] 999)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    27
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23152
diff changeset
    28
notation
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    29
  Rep_cfun  ("(_\<cdot>/_)" [999,1000] 999)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    30
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    31
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    32
subsection \<open>Syntax for continuous lambda abstraction\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    33
41479
655f583840d0 use proper syntactic types for 'syntax' commands
huffman
parents: 41478
diff changeset
    34
syntax "_cabs" :: "[logic, logic] \<Rightarrow> logic"
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
    35
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    36
parse_translation \<open>
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    37
(* rewrite (_cabs x t) => (Abs_cfun (%x. t)) *)
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42264
diff changeset
    38
  [Syntax_Trans.mk_binder_tr (@{syntax_const "_cabs"}, @{const_syntax Abs_cfun})];
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    39
\<close>
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    40
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    41
print_translation \<open>
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    42
  [(@{const_syntax Abs_cfun}, fn _ => fn [Abs abs] =>
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 42264
diff changeset
    43
      let val (x, t) = Syntax_Trans.atomic_abs_tr' abs
41478
18500bd1f47b make print translation for Abs_cfun consistent with other binders: prevent eta-contraction, but don't force eta-expansion
huffman
parents: 41430
diff changeset
    44
      in Syntax.const @{syntax_const "_cabs"} $ x $ t end)]
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    45
\<close>  \<comment> \<open>To avoid eta-contraction of body\<close>
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    46
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    47
text \<open>Syntax for nested abstractions\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    48
61998
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    49
syntax (ASCII)
41479
655f583840d0 use proper syntactic types for 'syntax' commands
huffman
parents: 41478
diff changeset
    50
  "_Lambda" :: "[cargs, logic] \<Rightarrow> logic"  ("(3LAM _./ _)" [1000, 10] 10)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    51
61998
b66d2ca1f907 clarified print modes;
wenzelm
parents: 61378
diff changeset
    52
syntax
41479
655f583840d0 use proper syntactic types for 'syntax' commands
huffman
parents: 41478
diff changeset
    53
  "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" ("(3\<Lambda> _./ _)" [1000, 10] 10)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
    54
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    55
parse_ast_translation \<open>
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    56
(* rewrite (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    57
(* 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
    58
  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
    59
    fun Lambda_ast_tr [pats, body] =
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    60
          Ast.fold_ast_p @{syntax_const "_cabs"}
42264
b6c1b0c4c511 separate structure Term_Position;
wenzelm
parents: 42224
diff changeset
    61
            (Ast.unfold_ast @{syntax_const "_cargs"} (Ast.strip_positions pats), body)
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    62
      | Lambda_ast_tr asts = raise Ast.AST ("Lambda_ast_tr", asts);
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    63
  in [(@{syntax_const "_Lambda"}, K Lambda_ast_tr)] end;
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    64
\<close>
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
    65
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    66
print_ast_translation \<open>
35115
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    67
(* rewrite (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)
446c5063e4fd modernized translations;
wenzelm
parents: 31076
diff changeset
    68
(* 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
    69
  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
    70
    fun cabs_ast_tr' asts =
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    71
      (case Ast.unfold_ast_p @{syntax_const "_cabs"}
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    72
          (Ast.Appl (Ast.Constant @{syntax_const "_cabs"} :: asts)) of
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    73
        ([], _) => raise Ast.AST ("cabs_ast_tr'", asts)
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    74
      | (xs, body) => Ast.Appl
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    75
          [Ast.Constant @{syntax_const "_Lambda"},
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42151
diff changeset
    76
           Ast.fold_ast @{syntax_const "_cargs"} xs, body]);
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51717
diff changeset
    77
  in [(@{syntax_const "_cabs"}, K cabs_ast_tr')] end
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    78
\<close>
15641
b389f108c485 added theorems eta_cfun and cont2cont_eta
huffman
parents: 15600
diff changeset
    79
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    80
text \<open>Dummy patterns for continuous abstraction\<close>
18079
9d4d70b175fd add translation for wildcard pattern
huffman
parents: 18078
diff changeset
    81
translations
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    82
  "\<Lambda> _. t" \<rightharpoonup> "CONST Abs_cfun (\<lambda>_. t)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    83
18087
577d57e51f89 add print translation: Abs_CFun f => LAM x. f x
huffman
parents: 18079
diff changeset
    84
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    85
subsection \<open>Continuous function space is pointed\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    86
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41400
diff changeset
    87
lemma bottom_cfun: "\<bottom> \<in> cfun"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    88
  by (simp add: cfun_def inst_fun_pcpo)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    89
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 35168
diff changeset
    90
instance cfun :: (cpo, discrete_cpo) discrete_cpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    91
  by intro_classes (simp add: below_cfun_def Rep_cfun_inject)
26025
ca6876116bb4 instances for class discrete_cpo
huffman
parents: 25927
diff changeset
    92
35525
fa231b86cb1e proper names for types cfun, sprod, ssum
huffman
parents: 35168
diff changeset
    93
instance cfun :: (cpo, pcpo) pcpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
    94
  by (rule typedef_pcpo [OF type_definition_cfun below_cfun_def bottom_cfun])
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
    95
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    96
lemmas Rep_cfun_strict =
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41400
diff changeset
    97
  typedef_Rep_strict [OF type_definition_cfun below_cfun_def bottom_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
    98
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
    99
lemmas Abs_cfun_strict =
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41400
diff changeset
   100
  typedef_Abs_strict [OF type_definition_cfun below_cfun_def bottom_cfun]
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   101
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   102
text \<open>function application is strict in its first argument\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   103
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   104
lemma Rep_cfun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   105
  by (simp add: Rep_cfun_strict)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   106
35641
a17bc4cec23a add simp rule LAM_strict
huffman
parents: 35547
diff changeset
   107
lemma LAM_strict [simp]: "(\<Lambda> x. \<bottom>) = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   108
  by (simp add: inst_fun_pcpo [symmetric] Abs_cfun_strict)
35641
a17bc4cec23a add simp rule LAM_strict
huffman
parents: 35547
diff changeset
   109
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   110
text \<open>for compatibility with old HOLCF-Version\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   111
lemma inst_cfun_pcpo: "\<bottom> = (\<Lambda> x. \<bottom>)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   112
  by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   113
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   114
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   115
subsection \<open>Basic properties of continuous functions\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   116
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   117
text \<open>Beta-equality for continuous functions\<close>
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
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   119
lemma Abs_cfun_inverse2: "cont f \<Longrightarrow> Rep_cfun (Abs_cfun f) = f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   120
  by (simp add: Abs_cfun_inverse cfun_def)
16098
6aef81a6ddd3 use TypedefPcpo for all class instances
huffman
parents: 16094
diff changeset
   121
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   122
lemma beta_cfun: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   123
  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
   124
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   125
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   126
subsubsection \<open>Beta-reduction simproc\<close>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   127
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   128
text \<open>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   129
  Given the term @{term "(\<Lambda> x. f x)\<cdot>y"}, the procedure tries to
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   130
  construct the theorem @{term "(\<Lambda> x. f x)\<cdot>y \<equiv> f y"}.  If this
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   131
  theorem cannot be completely solved by the cont2cont rules, then
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   132
  the procedure returns the ordinary conditional \<open>beta_cfun\<close>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   133
  rule.
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   134
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   135
  The simproc does not solve any more goals that would be solved by
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   136
  using \<open>beta_cfun\<close> as a simp rule.  The advantage of the
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   137
  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
   138
  that would otherwise be caused by large continuity side conditions.
41322
43a5b9a0ee8a beta-reduction simproc uses lemma Abs_cfun_inverse2 instead of beta_cfun, to be more robust against eta-contraction
huffman
parents: 41031
diff changeset
   139
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   140
  Update: The simproc now uses rule \<open>Abs_cfun_inverse2\<close> instead
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   141
  of \<open>beta_cfun\<close>, to avoid problems with eta-contraction.
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   142
\<close>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   143
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   144
simproc_setup beta_cfun_proc ("Rep_cfun (Abs_cfun f)") = \<open>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49759
diff changeset
   145
  fn phi => fn ctxt => fn ct =>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   146
    let
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   147
      val f = #2 (Thm.dest_comb (#2 (Thm.dest_comb ct)));
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   148
      val [T, U] = Thm.dest_ctyp (Thm.ctyp_of_cterm f);
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   149
      val tr = Thm.instantiate' [SOME T, SOME U] [SOME f] (mk_meta_eq @{thm Abs_cfun_inverse2});
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   150
      val rules = Named_Theorems.get ctxt \<^named_theorems>\<open>cont2cont\<close>;
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58880
diff changeset
   151
      val tac = SOLVED' (REPEAT_ALL_NEW (match_tac ctxt (rev rules)));
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   152
    in SOME (perhaps (SINGLE (tac 1)) tr) end
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   153
\<close>
37083
03a70ab79dd9 add beta_cfun simproc, which uses cont2cont rules
huffman
parents: 37079
diff changeset
   154
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   155
text \<open>Eta-equality for continuous functions\<close>
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
   156
36ee7f6af79f removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
huffman
parents: 16098
diff changeset
   157
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   158
  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
   159
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   160
text \<open>Extensionality for continuous functions\<close>
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
   161
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
   162
lemma cfun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f\<cdot>x = g\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   163
  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
   164
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
   165
lemma cfun_eqI: "(\<And>x. f\<cdot>x = g\<cdot>x) \<Longrightarrow> f = g"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   166
  by (simp add: cfun_eq_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   167
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   168
text \<open>Extensionality wrt. ordering for continuous functions\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   169
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   170
lemma cfun_below_iff: "f \<sqsubseteq> g \<longleftrightarrow> (\<forall>x. f\<cdot>x \<sqsubseteq> g\<cdot>x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   171
  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
   172
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
   173
lemma cfun_belowI: "(\<And>x. f\<cdot>x \<sqsubseteq> g\<cdot>x) \<Longrightarrow> f \<sqsubseteq> g"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   174
  by (simp add: cfun_below_iff)
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   175
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   176
text \<open>Congruence for continuous function application\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   177
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   178
lemma cfun_cong: "f = g \<Longrightarrow> x = y \<Longrightarrow> f\<cdot>x = g\<cdot>y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   179
  by simp
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   180
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
   181
lemma cfun_fun_cong: "f = g \<Longrightarrow> f\<cdot>x = g\<cdot>x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   182
  by simp
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   183
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
   184
lemma cfun_arg_cong: "x = y \<Longrightarrow> f\<cdot>x = f\<cdot>y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   185
  by simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   186
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   187
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   188
subsection \<open>Continuity of application\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   189
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   190
lemma cont_Rep_cfun1: "cont (\<lambda>f. f\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   191
  by (rule cont_Rep_cfun [OF cont_id, THEN cont2cont_fun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   192
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   193
lemma cont_Rep_cfun2: "cont (\<lambda>x. f\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   194
  using Rep_cfun [where x = f] by (simp add: cfun_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   195
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   196
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
   197
45606
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 42284
diff changeset
   198
lemmas monofun_Rep_cfun1 = cont_Rep_cfun1 [THEN cont2mono]
b1e1508643b1 eliminated obsolete "standard";
wenzelm
parents: 42284
diff changeset
   199
lemmas monofun_Rep_cfun2 = cont_Rep_cfun2 [THEN cont2mono]
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
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   201
text \<open>contlub, cont properties of @{term Rep_cfun} in each argument\<close>
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
   202
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   203
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   204
  by (rule cont_Rep_cfun2 [THEN cont2contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   205
27413
3154f3765cc7 replace lub (range Y) with (LUB i. Y i)
huffman
parents: 27274
diff changeset
   206
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   207
  by (rule cont_Rep_cfun1 [THEN cont2contlubE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   208
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   209
text \<open>monotonicity of application\<close>
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
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
lemma monofun_cfun_fun: "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   212
  by (simp add: cfun_below_iff)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   213
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
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   215
  by (rule monofun_Rep_cfun2 [THEN monofunE])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   216
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   217
lemma monofun_cfun: "f \<sqsubseteq> g \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   218
  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
   219
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   220
text \<open>ch2ch - rules for the type @{typ "'a \<rightarrow> 'b"}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   221
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
   222
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   223
  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
   224
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   225
lemma ch2ch_Rep_cfunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   226
  by (rule monofun_Rep_cfun2 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   227
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   228
lemma ch2ch_Rep_cfunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   229
  by (rule monofun_Rep_cfun1 [THEN ch2ch_monofun])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   230
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   231
lemma ch2ch_Rep_cfun [simp]: "chain F \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>(Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   232
  by (simp add: chain_def monofun_cfun)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   233
25884
a69e665b7df8 declare ch2ch_LAM [simp]
huffman
parents: 25827
diff changeset
   234
lemma ch2ch_LAM [simp]:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   235
  "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> (\<And>i. cont (\<lambda>x. S i x)) \<Longrightarrow> chain (\<lambda>i. \<Lambda> x. S i x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   236
  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
   237
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   238
text \<open>contlub, cont properties of @{term Rep_cfun} in both arguments\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   239
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   240
lemma lub_APP: "chain F \<Longrightarrow> chain Y \<Longrightarrow> (\<Squnion>i. F i\<cdot>(Y i)) = (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   241
  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
   242
41027
c599955d9806 add lemmas lub_APP, lub_LAM
huffman
parents: 40834
diff changeset
   243
lemma lub_LAM:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   244
  assumes "\<And>x. chain (\<lambda>i. F i x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   245
    and "\<And>i. cont (\<lambda>x. F i x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   246
  shows "(\<Squnion>i. \<Lambda> x. F i x) = (\<Lambda> x. \<Squnion>i. F i x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   247
  using assms by (simp add: lub_cfun lub_fun ch2ch_lambda)
18092
2c5d5da79a1e renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
huffman
parents: 18091
diff changeset
   248
41027
c599955d9806 add lemmas lub_APP, lub_LAM
huffman
parents: 40834
diff changeset
   249
lemmas lub_distribs = lub_APP lub_LAM
25901
bb178c8251e0 added lemmas lub_distribs
huffman
parents: 25884
diff changeset
   250
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   251
text \<open>strictness\<close>
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
   252
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 strictI: "f\<cdot>x = \<bottom> \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   254
  apply (rule bottomI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   255
  apply (erule subst)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   256
  apply (rule minimal [THEN monofun_cfun_arg])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   257
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   258
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   259
text \<open>type @{typ "'a \<rightarrow> 'b"} is chain complete\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   260
41031
9883d1417ce1 remove lemma cont_cfun;
huffman
parents: 41030
diff changeset
   261
lemma lub_cfun: "chain F \<Longrightarrow> (\<Squnion>i. F i) = (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   262
  by (simp add: lub_cfun lub_fun ch2ch_lambda)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   263
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   264
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   265
subsection \<open>Continuity simplification procedure\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   266
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   267
text \<open>cont2cont lemma for @{term Rep_cfun}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   268
40326
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40093
diff changeset
   269
lemma cont2cont_APP [simp, cont2cont]:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   270
  assumes f: "cont (\<lambda>x. f x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   271
  assumes t: "cont (\<lambda>x. t x)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   272
  shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   273
proof -
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   274
  from cont_Rep_cfun1 f have "cont (\<lambda>x. (f x)\<cdot>y)" for y
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   275
    by (rule cont_compose)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   276
  with t cont_Rep_cfun2 show "cont (\<lambda>x. (f x)\<cdot>(t x))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   277
    by (rule cont_apply)
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   278
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   279
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   280
text \<open>
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   281
  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
   282
  These lemmas are needed in theories that use types like @{typ "'a \<rightarrow> 'b \<Rightarrow> 'c"}.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   283
\<close>
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   284
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   285
lemma cont_APP_app [simp]: "cont f \<Longrightarrow> cont g \<Longrightarrow> cont (\<lambda>x. ((f x)\<cdot>(g x)) s)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   286
  by (rule cont2cont_APP [THEN cont2cont_fun])
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   287
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   288
lemma cont_APP_app_app [simp]: "cont f \<Longrightarrow> cont g \<Longrightarrow> cont (\<lambda>x. ((f x)\<cdot>(g x)) s t)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   289
  by (rule cont_APP_app [THEN cont2cont_fun])
40008
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   290
58ead6f77f8e move lemmas from Lift.thy to Cfun.thy
huffman
parents: 40006
diff changeset
   291
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   292
text \<open>cont2mono Lemma for @{term "\<lambda>x. LAM y. c1(x)(y)"}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   293
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   294
lemma cont2mono_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   295
  "\<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
   296
    \<Longrightarrow> monofun (\<lambda>x. \<Lambda> y. f x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   297
  by (simp add: monofun_def cfun_below_iff)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   298
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   299
text \<open>cont2cont Lemma for @{term "\<lambda>x. LAM y. f x y"}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   300
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   301
text \<open>
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   302
  Not suitable as a cont2cont rule, because on nested lambdas
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   303
  it causes exponential blow-up in the number of subgoals.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   304
\<close>
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   305
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   306
lemma cont2cont_LAM:
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   307
  assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   308
  assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   309
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   310
proof (rule cont_Abs_cfun)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   311
  from f1 show "f x \<in> cfun" for x
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   312
    by (simp add: cfun_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   313
  from f2 show "cont f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   314
    by (rule cont2cont_lambda)
29049
4e5b9e508e1e cleaned up some proofs in Cfun.thy
huffman
parents: 27413
diff changeset
   315
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   316
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   317
text \<open>
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   318
  This version does work as a cont2cont rule, since it
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   319
  has only a single subgoal.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   320
\<close>
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   321
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   322
lemma cont2cont_LAM' [simp, cont2cont]:
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   323
  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   324
  assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   325
  shows "cont (\<lambda>x. \<Lambda> y. f x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   326
  using assms by (simp add: cont2cont_LAM prod_cont_iff)
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   327
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   328
lemma cont2cont_LAM_discrete [simp, cont2cont]:
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   329
  "(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   330
  by (simp add: cont2cont_LAM)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   331
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   332
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   333
subsection \<open>Miscellaneous\<close>
17832
e18fc1a9a0e0 rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
huffman
parents: 17817
diff changeset
   334
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   335
text \<open>Monotonicity of @{term Abs_cfun}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   336
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   337
lemma monofun_LAM: "cont f \<Longrightarrow> cont g \<Longrightarrow> (\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> (\<Lambda> x. f x) \<sqsubseteq> (\<Lambda> x. g x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   338
  by (simp add: cfun_below_iff)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   339
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   340
text \<open>some lemmata for functions with flat/chfin domain/range types\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   341
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   342
lemma chfin_Rep_cfunR: "chain Y \<Longrightarrow> \<forall>s. \<exists>n. (LUB i. Y i)\<cdot>s = Y n\<cdot>s"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   343
  for Y :: "nat \<Rightarrow> 'a::cpo \<rightarrow> 'b::chfin"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   344
  apply (rule allI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   345
  apply (subst contlub_cfun_fun)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   346
   apply assumption
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   347
  apply (fast intro!: lub_eqI chfin lub_finch2 chfin2finch ch2ch_Rep_cfunL)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   348
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   349
18089
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   350
lemma adm_chfindom: "adm (\<lambda>(u::'a::cpo \<rightarrow> 'b::chfin). P(u\<cdot>s))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   351
  by (rule adm_subst, simp, rule adm_chfin)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   352
18089
35c091a9841a moved adm_chfindom from Fix.thy to Cfun.thy
huffman
parents: 18087
diff changeset
   353
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   354
subsection \<open>Continuous injection-retraction pairs\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   355
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   356
text \<open>Continuous retractions are strict.\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   357
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   358
lemma retraction_strict: "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   359
  apply (rule bottomI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   360
  apply (drule_tac x="\<bottom>" in spec)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   361
  apply (erule subst)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   362
  apply (rule monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   363
  apply (rule minimal)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   364
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   365
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   366
lemma injection_eq: "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x = g\<cdot>y) = (x = y)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   367
  apply (rule iffI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   368
   apply (drule_tac f=f in cfun_arg_cong)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   369
   apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   370
  apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   371
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   372
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   373
lemma injection_below: "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x \<sqsubseteq> g\<cdot>y) = (x \<sqsubseteq> y)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   374
  apply (rule iffI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   375
   apply (drule_tac f=f in monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   376
   apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   377
  apply (erule monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   378
  done
16314
7102a1aaecfd added theorem injection_less
huffman
parents: 16209
diff changeset
   379
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   380
lemma injection_defined_rev: "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> g\<cdot>z = \<bottom> \<Longrightarrow> z = \<bottom>"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   381
  apply (drule_tac f=f in cfun_arg_cong)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   382
  apply (simp add: retraction_strict)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   383
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   384
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   385
lemma injection_defined: "\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> z \<noteq> \<bottom> \<Longrightarrow> g\<cdot>z \<noteq> \<bottom>"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   386
  by (erule contrapos_nn, rule injection_defined_rev)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   387
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   388
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   389
text \<open>a result about functions with flat codomain\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   390
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   391
lemma flat_eqI: "x \<sqsubseteq> y \<Longrightarrow> x \<noteq> \<bottom> \<Longrightarrow> x = y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   392
  for x y :: "'a::flat"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   393
  by (drule ax_flat) simp
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   394
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   395
lemma flat_codom: "f\<cdot>x = c \<Longrightarrow> f\<cdot>\<bottom> = \<bottom> \<or> (\<forall>z. f\<cdot>z = c)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   396
  for c :: "'b::flat"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   397
  apply (case_tac "f\<cdot>x = \<bottom>")
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   398
   apply (rule disjI1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   399
   apply (rule bottomI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   400
   apply (erule_tac t="\<bottom>" in subst)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   401
   apply (rule minimal [THEN monofun_cfun_arg])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   402
  apply clarify
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   403
  apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   404
   apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   405
  apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   406
  done
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   407
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   408
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   409
subsection \<open>Identity and composition\<close>
15589
69bea57212ef reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   410
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   411
definition ID :: "'a \<rightarrow> 'a"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   412
  where "ID = (\<Lambda> x. x)"
25135
4f8176c940cf modernized specifications ('definition', 'axiomatization');
wenzelm
parents: 25131
diff changeset
   413
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   414
definition cfcomp  :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   415
  where 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
   416
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   417
abbreviation cfcomp_syn :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c"  (infixr "oo" 100)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   418
  where "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
   419
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   420
lemma ID1 [simp]: "ID\<cdot>x = x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   421
  by (simp add: ID_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   422
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   423
lemma cfcomp1: "(f oo g) = (\<Lambda> x. f\<cdot>(g\<cdot>x))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   424
  by (simp add: oo_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   425
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   426
lemma cfcomp2 [simp]: "(f oo g)\<cdot>x = f\<cdot>(g\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   427
  by (simp add: cfcomp1)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   428
27274
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   429
lemma cfcomp_LAM: "cont g \<Longrightarrow> f oo (\<Lambda> x. g x) = (\<Lambda> x. f\<cdot>(g x))"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   430
  by (simp add: cfcomp1)
27274
1c97c471db82 add lemma cfcomp_LAM
huffman
parents: 26025
diff changeset
   431
19709
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   432
lemma cfcomp_strict [simp]: "\<bottom> oo f = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   433
  by (simp add: cfun_eq_iff)
19709
78cd5f6af8e8 add theorem cfcomp_strict
huffman
parents: 18092
diff changeset
   434
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   435
text \<open>
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   436
  Show that interpretation of (pcpo, \<open>_\<rightarrow>_\<close>) is a category.
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   437
  \<^item> The class of objects is interpretation of syntactical class pcpo.
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   438
  \<^item> The class of arrows  between objects @{typ 'a} and @{typ 'b} is interpret. of @{typ "'a \<rightarrow> 'b"}.
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   439
  \<^item> The identity arrow is interpretation of @{term ID}.
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   440
  \<^item> The composition of f and g is interpretation of \<open>oo\<close>.
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   441
\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   442
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   443
lemma ID2 [simp]: "f oo ID = f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   444
  by (rule cfun_eqI, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   445
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   446
lemma ID3 [simp]: "ID oo f = f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   447
  by (rule cfun_eqI) simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   448
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   449
lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   450
  by (rule cfun_eqI) simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   451
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   452
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   453
subsection \<open>Strictified functions\<close>
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   454
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35933
diff changeset
   455
default_sort pcpo
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   456
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   457
definition seq :: "'a \<rightarrow> 'b \<rightarrow> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   458
  where "seq = (\<Lambda> x. if x = \<bottom> then \<bottom> else ID)"
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   459
40794
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   460
lemma cont2cont_if_bottom [cont2cont, simp]:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   461
  assumes f: "cont (\<lambda>x. f x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   462
    and g: "cont (\<lambda>x. g x)"
40794
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   463
  shows "cont (\<lambda>x. if f x = \<bottom> then \<bottom> else g x)"
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   464
proof (rule cont_apply [OF f])
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   465
  show "cont (\<lambda>y. if y = \<bottom> then \<bottom> else g x)" for x
40794
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   466
    unfolding cont_def is_lub_def is_ub_def ball_simps
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   467
    by (simp add: lub_eq_bottom_iff)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   468
  show "cont (\<lambda>x. if y = \<bottom> then \<bottom> else g x)" for y
40794
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   469
    by (simp add: g)
d28d41ee4cef add lemma cont2cont_if_bottom
huffman
parents: 40774
diff changeset
   470
qed
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   471
40767
a3e505b236e7 rename function 'strict' to 'seq', which is its name in Haskell
huffman
parents: 40502
diff changeset
   472
lemma seq_conv_if: "seq\<cdot>x = (if x = \<bottom> then \<bottom> else ID)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   473
  by (simp add: seq_def)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   474
41400
1a7557cc686a replaced separate lemmas seq{1,2,3} with seq_simps
huffman
parents: 41322
diff changeset
   475
lemma seq_simps [simp]:
1a7557cc686a replaced separate lemmas seq{1,2,3} with seq_simps
huffman
parents: 41322
diff changeset
   476
  "seq\<cdot>\<bottom> = \<bottom>"
1a7557cc686a replaced separate lemmas seq{1,2,3} with seq_simps
huffman
parents: 41322
diff changeset
   477
  "seq\<cdot>x\<cdot>\<bottom> = \<bottom>"
1a7557cc686a replaced separate lemmas seq{1,2,3} with seq_simps
huffman
parents: 41322
diff changeset
   478
  "x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x = ID"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   479
  by (simp_all add: seq_conv_if)
40093
c2d36bc4cd14 add lemma strict3
huffman
parents: 40091
diff changeset
   480
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   481
definition strictify  :: "('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   482
  where "strictify = (\<Lambda> f x. seq\<cdot>x\<cdot>(f\<cdot>x))"
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   483
17815
ccf54e3cabfa removed Istrictify; simplified some proofs
huffman
parents: 16920
diff changeset
   484
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   485
  by (simp add: strictify_def)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   486
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   487
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   488
  by (simp add: strictify_conv_if)
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   489
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   490
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   491
  by (simp add: strictify_conv_if)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   492
16085
c004b9bc970e rewrote continuous isomorphism section, cleaned up
huffman
parents: 16070
diff changeset
   493
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   494
subsection \<open>Continuity of let-bindings\<close>
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   495
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   496
lemma cont2cont_Let:
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   497
  assumes f: "cont (\<lambda>x. f x)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   498
  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
   499
  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
   500
  shows "cont (\<lambda>x. let y = f x in g x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   501
  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
   502
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36452
diff changeset
   503
lemma cont2cont_Let' [simp, cont2cont]:
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   504
  assumes f: "cont (\<lambda>x. f x)"
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   505
  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
   506
  shows "cont (\<lambda>x. let y = f x in g x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   507
  using f
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   508
proof (rule cont2cont_Let)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   509
  from g show "cont (\<lambda>y. g x y)" for x
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   510
    by (simp add: prod_cont_iff)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   511
  from g show "cont (\<lambda>x. g x y)" for y
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   512
    by (simp add: prod_cont_iff)
35933
f135ebcc835c remove continuous let-binding function CLet; add cont2cont rule ordinary Let
huffman
parents: 35914
diff changeset
   513
qed
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17815
diff changeset
   514
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   515
text \<open>The simple version (suggested by Joachim Breitner) is needed if
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   516
  the type of the defined term is not a cpo.\<close>
39145
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   517
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   518
lemma cont2cont_Let_simple [simp, cont2cont]:
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   519
  assumes "\<And>y. cont (\<lambda>x. g x y)"
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   520
  shows "cont (\<lambda>x. let y = t in g x y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63549
diff changeset
   521
  unfolding Let_def using assms .
39145
154fd9c06c63 add lemma cont2cont_Let_simple
huffman
parents: 37083
diff changeset
   522
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   523
end