| author | wenzelm | 
| Wed, 09 Apr 2014 12:33:02 +0200 | |
| changeset 56493 | 1f660d858a75 | 
| parent 52143 | 36ffe23b25f8 | 
| child 57945 | cacb00a569e0 | 
| permissions | -rw-r--r-- | 
| 42151 | 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 | 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 | 8  | 
theory Cfun  | 
| 40772 | 9  | 
imports Cpodef Fun_Cpo Product_Cpo  | 
| 15577 | 10  | 
begin  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
11  | 
|
| 36452 | 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  | 
|
| 45695 | 16  | 
definition "cfun = {f::'a => 'b. cont f}"
 | 
17  | 
||
| 
49759
 
ecf1d5d87e5e
removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
 
huffman 
parents: 
45695 
diff
changeset
 | 
18  | 
cpodef ('a, 'b) cfun (infixr "->" 0) = "cfun :: ('a => 'b) set"
 | 
| 45695 | 19  | 
unfolding cfun_def by (auto intro: cont_const adm_cont)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
20  | 
|
| 35427 | 21  | 
type_notation (xsymbols)  | 
| 35525 | 22  | 
  cfun  ("(_ \<rightarrow>/ _)" [1, 0] 0)
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
23  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
24  | 
notation  | 
| 40327 | 25  | 
  Rep_cfun  ("(_$/_)" [999,1000] 999)
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
26  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
27  | 
notation (xsymbols)  | 
| 40327 | 28  | 
  Rep_cfun  ("(_\<cdot>/_)" [999,1000] 999)
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
29  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
30  | 
notation (HTML output)  | 
| 40327 | 31  | 
  Rep_cfun  ("(_\<cdot>/_)" [999,1000] 999)
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
32  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
33  | 
subsection {* Syntax for continuous lambda abstraction *}
 | 
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
34  | 
|
| 41479 | 35  | 
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
 | 
36  | 
|
| 
 
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
 | 
37  | 
parse_translation {*
 | 
| 40327 | 38  | 
(* rewrite (_cabs x t) => (Abs_cfun (%x. t)) *)  | 
| 42284 | 39  | 
  [Syntax_Trans.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
 | 
40  | 
*}  | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
41  | 
|
| 
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
 | 
42  | 
print_translation {*
 | 
| 52143 | 43  | 
  [(@{const_syntax Abs_cfun}, fn _ => fn [Abs abs] =>
 | 
| 42284 | 44  | 
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
 | 
45  | 
      in Syntax.const @{syntax_const "_cabs"} $ x $ t end)]
 | 
| 
 
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
 | 
46  | 
*}  -- {* To avoid eta-contraction of body *}
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
47  | 
|
| 18087 | 48  | 
text {* Syntax for nested abstractions *}
 | 
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
49  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
50  | 
syntax  | 
| 41479 | 51  | 
  "_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
 | 
52  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
53  | 
syntax (xsymbols)  | 
| 41479 | 54  | 
  "_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
 | 
55  | 
|
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
56  | 
parse_ast_translation {*
 | 
| 35115 | 57  | 
(* rewrite (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)  | 
58  | 
(* 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
 | 
59  | 
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
 | 
60  | 
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
 | 
61  | 
          Ast.fold_ast_p @{syntax_const "_cabs"}
 | 
| 42264 | 62  | 
            (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
 | 
63  | 
      | Lambda_ast_tr asts = raise Ast.AST ("Lambda_ast_tr", asts);
 | 
| 52143 | 64  | 
  in [(@{syntax_const "_Lambda"}, K Lambda_ast_tr)] end;
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
65  | 
*}  | 
| 
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
66  | 
|
| 
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
67  | 
print_ast_translation {*
 | 
| 35115 | 68  | 
(* rewrite (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)  | 
69  | 
(* 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
 | 
70  | 
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
 | 
71  | 
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
 | 
72  | 
      (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
 | 
73  | 
          (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
 | 
74  | 
        ([], _) => 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
 | 
75  | 
| (xs, body) => Ast.Appl  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
76  | 
          [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
 | 
77  | 
           Ast.fold_ast @{syntax_const "_cargs"} xs, body]);
 | 
| 52143 | 78  | 
  in [(@{syntax_const "_cabs"}, K cabs_ast_tr')] end
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
79  | 
*}  | 
| 15641 | 80  | 
|
| 18087 | 81  | 
text {* Dummy patterns for continuous abstraction *}
 | 
| 18079 | 82  | 
translations  | 
| 40327 | 83  | 
"\<Lambda> _. t" => "CONST Abs_cfun (\<lambda> _. t)"  | 
| 18087 | 84  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
85  | 
subsection {* Continuous function space is pointed *}
 | 
| 
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"  | 
| 40327 | 88  | 
by (simp add: cfun_def inst_fun_pcpo)  | 
| 16098 | 89  | 
|
| 35525 | 90  | 
instance cfun :: (cpo, discrete_cpo) discrete_cpo  | 
| 40327 | 91  | 
by intro_classes (simp add: below_cfun_def Rep_cfun_inject)  | 
| 26025 | 92  | 
|
| 35525 | 93  | 
instance cfun :: (cpo, pcpo) pcpo  | 
| 
41430
 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 
huffman 
parents: 
41400 
diff
changeset
 | 
94  | 
by (rule typedef_pcpo [OF type_definition_cfun below_cfun_def bottom_cfun])  | 
| 16098 | 95  | 
|
| 40327 | 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 | 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 | 101  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
102  | 
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
 | 
103  | 
|
| 40327 | 104  | 
lemma Rep_cfun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"  | 
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 | 107  | 
lemma LAM_strict [simp]: "(\<Lambda> x. \<bottom>) = \<bottom>"  | 
| 40327 | 108  | 
by (simp add: inst_fun_pcpo [symmetric] Abs_cfun_strict)  | 
| 35641 | 109  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
110  | 
text {* for compatibility with old HOLCF-Version *}
 | 
| 
 
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>)"  | 
| 35641 | 112  | 
by simp  | 
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
113  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
114  | 
subsection {* Basic properties of continuous functions *}
 | 
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
115  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
116  | 
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
 | 
117  | 
|
| 40327 | 118  | 
lemma Abs_cfun_inverse2: "cont f \<Longrightarrow> Rep_cfun (Abs_cfun f) = f"  | 
119  | 
by (simp add: Abs_cfun_inverse cfun_def)  | 
|
| 16098 | 120  | 
|
| 
37083
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
121  | 
lemma beta_cfun: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"  | 
| 40327 | 122  | 
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
 | 
123  | 
|
| 
37083
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
124  | 
text {* Beta-reduction simproc *}
 | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
125  | 
|
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
126  | 
text {*
 | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
127  | 
  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
 | 
128  | 
  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
 | 
129  | 
theorem cannot be completely solved by the cont2cont rules, then  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
130  | 
  the procedure returns the ordinary conditional @{text beta_cfun}
 | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
131  | 
rule.  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
132  | 
|
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
133  | 
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
 | 
134  | 
  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
 | 
135  | 
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
 | 
136  | 
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
 | 
137  | 
|
| 
 
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
 | 
138  | 
  Update: The simproc now uses rule @{text Abs_cfun_inverse2} instead
 | 
| 
 
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  | 
  of @{text beta_cfun}, to avoid problems with eta-contraction.
 | 
| 
37083
 
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  | 
|
| 
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
 | 
142  | 
simproc_setup beta_cfun_proc ("Rep_cfun (Abs_cfun f)") = {*
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49759 
diff
changeset
 | 
143  | 
fn phi => fn ctxt => fn ct =>  | 
| 
37083
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
144  | 
let  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
145  | 
val dest = Thm.dest_comb;  | 
| 
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
 | 
146  | 
val f = (snd o dest o snd o dest) ct;  | 
| 
37083
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
147  | 
val [T, U] = Thm.dest_ctyp (ctyp_of_term f);  | 
| 
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
 | 
148  | 
val tr = instantiate' [SOME T, SOME U] [SOME f]  | 
| 
 
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
 | 
149  | 
          (mk_meta_eq @{thm Abs_cfun_inverse2});
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
49759 
diff
changeset
 | 
150  | 
val rules = Cont2ContData.get ctxt;  | 
| 
37083
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
151  | 
val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
152  | 
in SOME (perhaps (SINGLE (tac 1)) tr) end  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
153  | 
*}  | 
| 
 
03a70ab79dd9
add beta_cfun simproc, which uses cont2cont rules
 
huffman 
parents: 
37079 
diff
changeset
 | 
154  | 
|
| 
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
 | 
155  | 
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
 | 
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"  | 
| 40327 | 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  | 
|
| 
 
36ee7f6af79f
removed 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 {* 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
 | 
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)"  | 
| 40327 | 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"  | 
| 
 
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
 
huffman 
parents: 
40001 
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  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
168  | 
text {* Extensionality wrt. ordering for continuous functions *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
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_below_iff: "f \<sqsubseteq> g \<longleftrightarrow> (\<forall>x. f\<cdot>x \<sqsubseteq> g\<cdot>x)"  | 
| 40327 | 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"  | 
| 
 
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
 
huffman 
parents: 
40001 
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  | 
|
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
176  | 
text {* Congruence for continuous function application *}
 | 
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
177  | 
|
| 
16209
 
36ee7f6af79f
removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
 
huffman 
parents: 
16098 
diff
changeset
 | 
178  | 
lemma 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
 | 
179  | 
by simp  | 
| 
 
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"  | 
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
182  | 
by simp  | 
| 
 
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"  | 
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
185  | 
by simp  | 
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
186  | 
|
| 
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
 | 
187  | 
subsection {* Continuity of application *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
188  | 
|
| 40327 | 189  | 
lemma cont_Rep_cfun1: "cont (\<lambda>f. f\<cdot>x)"  | 
| 
40834
 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 
huffman 
parents: 
40794 
diff
changeset
 | 
190  | 
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
 | 
191  | 
|
| 40327 | 192  | 
lemma cont_Rep_cfun2: "cont (\<lambda>x. f\<cdot>x)"  | 
193  | 
apply (cut_tac x=f in Rep_cfun)  | 
|
194  | 
apply (simp add: cfun_def)  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
195  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
196  | 
|
| 40327 | 197  | 
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
 | 
198  | 
|
| 45606 | 199  | 
lemmas monofun_Rep_cfun1 = cont_Rep_cfun1 [THEN cont2mono]  | 
200  | 
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
 | 
201  | 
|
| 40327 | 202  | 
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
 | 
203  | 
|
| 27413 | 204  | 
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"  | 
| 40327 | 205  | 
by (rule cont_Rep_cfun2 [THEN cont2contlubE])  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
206  | 
|
| 27413 | 207  | 
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"  | 
| 40327 | 208  | 
by (rule cont_Rep_cfun1 [THEN cont2contlubE])  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
209  | 
|
| 
16209
 
36ee7f6af79f
removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
 
huffman 
parents: 
16098 
diff
changeset
 | 
210  | 
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
 | 
211  | 
|
| 
 
36ee7f6af79f
removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
 
huffman 
parents: 
16098 
diff
changeset
 | 
212  | 
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
 | 
213  | 
by (simp add: cfun_below_iff)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
214  | 
|
| 
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
 | 
215  | 
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"  | 
| 40327 | 216  | 
by (rule monofun_Rep_cfun2 [THEN monofunE])  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
217  | 
|
| 
16209
 
36ee7f6af79f
removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
 
huffman 
parents: 
16098 
diff
changeset
 | 
218  | 
lemma 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
 | 
219  | 
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
 | 
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 {* ch2ch - rules for the type @{typ "'a -> 'b"} *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
222  | 
|
| 
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
 | 
223  | 
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"  | 
| 40327 | 224  | 
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
 | 
225  | 
|
| 40327 | 226  | 
lemma ch2ch_Rep_cfunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"  | 
227  | 
by (rule monofun_Rep_cfun2 [THEN ch2ch_monofun])  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
228  | 
|
| 40327 | 229  | 
lemma ch2ch_Rep_cfunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"  | 
230  | 
by (rule monofun_Rep_cfun1 [THEN ch2ch_monofun])  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
231  | 
|
| 40327 | 232  | 
lemma ch2ch_Rep_cfun [simp]:  | 
| 18076 | 233  | 
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>(Y i))"  | 
| 25884 | 234  | 
by (simp add: chain_def monofun_cfun)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
235  | 
|
| 25884 | 236  | 
lemma ch2ch_LAM [simp]:  | 
237  | 
"\<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
 | 
238  | 
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
 | 
239  | 
|
| 40327 | 240  | 
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
 | 
241  | 
|
| 41027 | 242  | 
lemma lub_APP:  | 
243  | 
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>i. F i\<cdot>(Y i)) = (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i)"  | 
|
| 18076 | 244  | 
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
 | 
245  | 
|
| 41027 | 246  | 
lemma lub_LAM:  | 
| 
18092
 
2c5d5da79a1e
renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
 
huffman 
parents: 
18091 
diff
changeset
 | 
247  | 
"\<lbrakk>\<And>x. chain (\<lambda>i. F i x); \<And>i. cont (\<lambda>x. F i x)\<rbrakk>  | 
| 41027 | 248  | 
\<Longrightarrow> (\<Squnion>i. \<Lambda> x. F i x) = (\<Lambda> x. \<Squnion>i. F i x)"  | 
| 
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
 | 
249  | 
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
 | 
250  | 
|
| 41027 | 251  | 
lemmas lub_distribs = lub_APP lub_LAM  | 
| 25901 | 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  | 
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
 | 
254  | 
|
| 
 
36ee7f6af79f
removed dependencies on MF2 lemmas; removed some obsolete theorems; cleaned up many proofs; renamed less_cfun2 to less_cfun_ext
 
huffman 
parents: 
16098 
diff
changeset
 | 
255  | 
lemma strictI: "f\<cdot>x = \<bottom> \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"  | 
| 
41430
 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 
huffman 
parents: 
41400 
diff
changeset
 | 
256  | 
apply (rule bottomI)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
257  | 
apply (erule subst)  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
258  | 
apply (rule minimal [THEN monofun_cfun_arg])  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
259  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
260  | 
|
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
261  | 
text {* type @{typ "'a -> 'b"} is chain complete *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
262  | 
|
| 41031 | 263  | 
lemma lub_cfun: "chain F \<Longrightarrow> (\<Squnion>i. F i) = (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"  | 
264  | 
by (simp add: lub_cfun lub_fun ch2ch_lambda)  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
265  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
266  | 
subsection {* Continuity simplification procedure *}
 | 
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
267  | 
|
| 40327 | 268  | 
text {* cont2cont lemma for @{term Rep_cfun} *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
269  | 
|
| 
40326
 
73d45866dbda
renamed lemma cont2cont_Rep_CFun to cont2cont_APP
 
huffman 
parents: 
40093 
diff
changeset
 | 
270  | 
lemma cont2cont_APP [simp, cont2cont]:  | 
| 29049 | 271  | 
assumes f: "cont (\<lambda>x. f x)"  | 
272  | 
assumes t: "cont (\<lambda>x. t x)"  | 
|
273  | 
shows "cont (\<lambda>x. (f x)\<cdot>(t x))"  | 
|
274  | 
proof -  | 
|
| 40006 | 275  | 
have 1: "\<And>y. cont (\<lambda>x. (f x)\<cdot>y)"  | 
| 40327 | 276  | 
using cont_Rep_cfun1 f by (rule cont_compose)  | 
| 40006 | 277  | 
show "cont (\<lambda>x. (f x)\<cdot>(t x))"  | 
| 40327 | 278  | 
using t cont_Rep_cfun2 1 by (rule cont_apply)  | 
| 29049 | 279  | 
qed  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
280  | 
|
| 40008 | 281  | 
text {*
 | 
282  | 
Two specific lemmas for the combination of LCF and HOL terms.  | 
|
283  | 
  These lemmas are needed in theories that use types like @{typ "'a \<rightarrow> 'b \<Rightarrow> 'c"}.
 | 
|
284  | 
*}  | 
|
285  | 
||
| 
40326
 
73d45866dbda
renamed lemma cont2cont_Rep_CFun to cont2cont_APP
 
huffman 
parents: 
40093 
diff
changeset
 | 
286  | 
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
 | 
287  | 
by (rule cont2cont_APP [THEN cont2cont_fun])  | 
| 40008 | 288  | 
|
| 
40326
 
73d45866dbda
renamed lemma cont2cont_Rep_CFun to cont2cont_APP
 
huffman 
parents: 
40093 
diff
changeset
 | 
289  | 
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
 | 
290  | 
by (rule cont_APP_app [THEN cont2cont_fun])  | 
| 40008 | 291  | 
|
292  | 
||
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
293  | 
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
 | 
294  | 
|
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
295  | 
lemma cont2mono_LAM:  | 
| 29049 | 296  | 
"\<lbrakk>\<And>x. cont (\<lambda>y. f x y); \<And>y. monofun (\<lambda>x. f x y)\<rbrakk>  | 
297  | 
\<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
 | 
298  | 
unfolding monofun_def cfun_below_iff by simp  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
299  | 
|
| 29049 | 300  | 
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
 | 
301  | 
|
| 
29530
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
302  | 
text {*
 | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
303  | 
Not suitable as a cont2cont rule, because on nested lambdas  | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
304  | 
it causes exponential blow-up in the number of subgoals.  | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
305  | 
*}  | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
306  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
307  | 
lemma cont2cont_LAM:  | 
| 29049 | 308  | 
assumes f1: "\<And>x. cont (\<lambda>y. f x y)"  | 
309  | 
assumes f2: "\<And>y. cont (\<lambda>x. f x y)"  | 
|
310  | 
shows "cont (\<lambda>x. \<Lambda> y. f x y)"  | 
|
| 40327 | 311  | 
proof (rule cont_Abs_cfun)  | 
| 29049 | 312  | 
fix x  | 
| 40327 | 313  | 
from f1 show "f x \<in> cfun" by (simp add: cfun_def)  | 
| 29049 | 314  | 
from f2 show "cont f" by (rule cont2cont_lambda)  | 
315  | 
qed  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
316  | 
|
| 
29530
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
317  | 
text {*
 | 
| 
 
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.  | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
320  | 
*}  | 
| 
 
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)"  | 
| 
39808
 
1410c84013b9
rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
 
huffman 
parents: 
39302 
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)"  | 
| 
 
9905b660612b
change to simpler, more extensible continuity simproc
 
huffman 
parents: 
29138 
diff
changeset
 | 
330  | 
by (simp add: cont2cont_LAM)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
331  | 
|
| 
17832
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
332  | 
subsection {* Miscellaneous *}
 | 
| 
 
e18fc1a9a0e0
rearranged subsections; added theorems expand_cfun_eq, expand_cfun_less
 
huffman 
parents: 
17817 
diff
changeset
 | 
333  | 
|
| 40327 | 334  | 
text {* Monotonicity of @{term Abs_cfun} *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
335  | 
|
| 
40433
 
3128c2a54785
remove some unnecessary lemmas; move monofun_LAM to Cfun.thy
 
huffman 
parents: 
40327 
diff
changeset
 | 
336  | 
lemma monofun_LAM:  | 
| 
 
3128c2a54785
remove some unnecessary lemmas; move monofun_LAM to Cfun.thy
 
huffman 
parents: 
40327 
diff
changeset
 | 
337  | 
"\<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
 | 
338  | 
by (simp add: cfun_below_iff)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
339  | 
|
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
340  | 
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
 | 
341  | 
|
| 40327 | 342  | 
lemma chfin_Rep_cfunR: "chain (Y::nat => 'a::cpo->'b::chfin)  | 
| 27413 | 343  | 
==> !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
 | 
344  | 
apply (rule allI)  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
345  | 
apply (subst contlub_cfun_fun)  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
346  | 
apply assumption  | 
| 40771 | 347  | 
apply (fast intro!: lub_eqI chfin lub_finch2 chfin2finch ch2ch_Rep_cfunL)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
348  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
349  | 
|
| 18089 | 350  | 
lemma adm_chfindom: "adm (\<lambda>(u::'a::cpo \<rightarrow> 'b::chfin). P(u\<cdot>s))"  | 
351  | 
by (rule adm_subst, simp, rule adm_chfin)  | 
|
352  | 
||
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
353  | 
subsection {* Continuous injection-retraction pairs *}
 | 
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
354  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
355  | 
text {* Continuous retractions are strict. *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
356  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
357  | 
lemma retraction_strict:  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
358  | 
"\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"  | 
| 
41430
 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 
huffman 
parents: 
41400 
diff
changeset
 | 
359  | 
apply (rule bottomI)  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
360  | 
apply (drule_tac x="\<bottom>" in spec)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
361  | 
apply (erule subst)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
362  | 
apply (rule monofun_cfun_arg)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
363  | 
apply (rule minimal)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
364  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
365  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
366  | 
lemma injection_eq:  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
367  | 
"\<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
 | 
368  | 
apply (rule iffI)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
369  | 
apply (drule_tac f=f in cfun_arg_cong)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
370  | 
apply simp  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
371  | 
apply simp  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
372  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
373  | 
|
| 
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
 | 
374  | 
lemma injection_below:  | 
| 16314 | 375  | 
"\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x \<sqsubseteq> g\<cdot>y) = (x \<sqsubseteq> y)"  | 
376  | 
apply (rule iffI)  | 
|
377  | 
apply (drule_tac f=f in monofun_cfun_arg)  | 
|
378  | 
apply simp  | 
|
379  | 
apply (erule monofun_cfun_arg)  | 
|
380  | 
done  | 
|
381  | 
||
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
382  | 
lemma injection_defined_rev:  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
383  | 
"\<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
 | 
384  | 
apply (drule_tac f=f in cfun_arg_cong)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
385  | 
apply (simp add: retraction_strict)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
386  | 
done  | 
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
387  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
388  | 
lemma injection_defined:  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
389  | 
"\<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
 | 
390  | 
by (erule contrapos_nn, rule injection_defined_rev)  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
391  | 
|
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
392  | 
text {* a result about functions with flat codomain *}
 | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
393  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
394  | 
lemma flat_eqI: "\<lbrakk>(x::'a::flat) \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> x = y"  | 
| 25920 | 395  | 
by (drule ax_flat, simp)  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
396  | 
|
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
397  | 
lemma flat_codom:  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
398  | 
"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
 | 
399  | 
apply (case_tac "f\<cdot>x = \<bottom>")  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
400  | 
apply (rule disjI1)  | 
| 
41430
 
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
 
huffman 
parents: 
41400 
diff
changeset
 | 
401  | 
apply (rule bottomI)  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
402  | 
apply (erule_tac t="\<bottom>" in subst)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
403  | 
apply (rule minimal [THEN monofun_cfun_arg])  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
404  | 
apply clarify  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
405  | 
apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
406  | 
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
407  | 
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
 | 
408  | 
done  | 
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
409  | 
|
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
410  | 
subsection {* Identity and composition *}
 | 
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
411  | 
|
| 
25135
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
412  | 
definition  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
413  | 
ID :: "'a \<rightarrow> 'a" where  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
414  | 
"ID = (\<Lambda> x. x)"  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
415  | 
|
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
416  | 
definition  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
417  | 
  cfcomp  :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c" where
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
25131 
diff
changeset
 | 
418  | 
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
 | 
419  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
420  | 
abbreviation  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
421  | 
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
 | 
422  | 
"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
 | 
423  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
424  | 
lemma ID1 [simp]: "ID\<cdot>x = x"  | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
425  | 
by (simp add: ID_def)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
426  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
427  | 
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
 | 
428  | 
by (simp add: oo_def)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
429  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
430  | 
lemma 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
 | 
431  | 
by (simp add: cfcomp1)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
432  | 
|
| 27274 | 433  | 
lemma cfcomp_LAM: "cont g \<Longrightarrow> f oo (\<Lambda> x. g x) = (\<Lambda> x. f\<cdot>(g x))"  | 
434  | 
by (simp add: cfcomp1)  | 
|
435  | 
||
| 19709 | 436  | 
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
 | 
437  | 
by (simp add: cfun_eq_iff)  | 
| 19709 | 438  | 
|
| 
15589
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
439  | 
text {*
 | 
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
440  | 
  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
 | 
441  | 
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
 | 
442  | 
  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
 | 
443  | 
  The identity arrow is interpretation of @{term ID}.
 | 
| 
 
69bea57212ef
reordered and arranged for document generation, cleaned up some proofs
 
huffman 
parents: 
15577 
diff
changeset
 | 
444  | 
  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
 | 
445  | 
*}  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
446  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
447  | 
lemma 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
 | 
448  | 
by (rule cfun_eqI, simp)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
449  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
450  | 
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
 | 
451  | 
by (rule cfun_eqI, simp)  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
452  | 
|
| 
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
453  | 
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
 | 
454  | 
by (rule cfun_eqI, simp)  | 
| 
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  | 
subsection {* Strictified functions *}
 | 
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
457  | 
|
| 36452 | 458  | 
default_sort pcpo  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
459  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
23152 
diff
changeset
 | 
460  | 
definition  | 
| 
40767
 
a3e505b236e7
rename function 'strict' to 'seq', which is its name in Haskell
 
huffman 
parents: 
40502 
diff
changeset
 | 
461  | 
seq :: "'a \<rightarrow> 'b \<rightarrow> 'b" where  | 
| 
 
a3e505b236e7
rename function 'strict' to 'seq', which is its name in Haskell
 
huffman 
parents: 
40502 
diff
changeset
 | 
462  | 
"seq = (\<Lambda> x. if x = \<bottom> then \<bottom> else ID)"  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
463  | 
|
| 40794 | 464  | 
lemma cont2cont_if_bottom [cont2cont, simp]:  | 
465  | 
assumes f: "cont (\<lambda>x. f x)" and g: "cont (\<lambda>x. g x)"  | 
|
466  | 
shows "cont (\<lambda>x. if f x = \<bottom> then \<bottom> else g x)"  | 
|
467  | 
proof (rule cont_apply [OF f])  | 
|
468  | 
show "\<And>x. cont (\<lambda>y. if y = \<bottom> then \<bottom> else g x)"  | 
|
469  | 
unfolding cont_def is_lub_def is_ub_def ball_simps  | 
|
470  | 
by (simp add: lub_eq_bottom_iff)  | 
|
471  | 
show "\<And>y. cont (\<lambda>x. if y = \<bottom> then \<bottom> else g x)"  | 
|
472  | 
by (simp add: g)  | 
|
473  | 
qed  | 
|
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
474  | 
|
| 
40767
 
a3e505b236e7
rename function 'strict' to 'seq', which is its name in Haskell
 
huffman 
parents: 
40502 
diff
changeset
 | 
475  | 
lemma seq_conv_if: "seq\<cdot>x = (if x = \<bottom> then \<bottom> else ID)"  | 
| 40794 | 476  | 
unfolding seq_def by simp  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
477  | 
|
| 
41400
 
1a7557cc686a
replaced separate lemmas seq{1,2,3} with seq_simps
 
huffman 
parents: 
41322 
diff
changeset
 | 
478  | 
lemma seq_simps [simp]:  | 
| 
 
1a7557cc686a
replaced separate lemmas seq{1,2,3} with seq_simps
 
huffman 
parents: 
41322 
diff
changeset
 | 
479  | 
"seq\<cdot>\<bottom> = \<bottom>"  | 
| 
 
1a7557cc686a
replaced separate lemmas seq{1,2,3} with seq_simps
 
huffman 
parents: 
41322 
diff
changeset
 | 
480  | 
"seq\<cdot>x\<cdot>\<bottom> = \<bottom>"  | 
| 
 
1a7557cc686a
replaced separate lemmas seq{1,2,3} with seq_simps
 
huffman 
parents: 
41322 
diff
changeset
 | 
481  | 
"x \<noteq> \<bottom> \<Longrightarrow> seq\<cdot>x = ID"  | 
| 
 
1a7557cc686a
replaced separate lemmas seq{1,2,3} with seq_simps
 
huffman 
parents: 
41322 
diff
changeset
 | 
482  | 
by (simp_all add: seq_conv_if)  | 
| 40093 | 483  | 
|
484  | 
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
 | 
485  | 
  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
 | 
486  | 
"strictify = (\<Lambda> f x. seq\<cdot>x\<cdot>(f\<cdot>x))"  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
487  | 
|
| 17815 | 488  | 
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
 | 
489  | 
unfolding strictify_def by simp  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
490  | 
|
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
491  | 
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"  | 
| 17815 | 492  | 
by (simp add: strictify_conv_if)  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
493  | 
|
| 
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
494  | 
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"  | 
| 17815 | 495  | 
by (simp add: strictify_conv_if)  | 
| 
16085
 
c004b9bc970e
rewrote continuous isomorphism section, cleaned up
 
huffman 
parents: 
16070 
diff
changeset
 | 
496  | 
|
| 
35933
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
497  | 
subsection {* Continuity of let-bindings *}
 | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
498  | 
|
| 
35933
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
499  | 
lemma cont2cont_Let:  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
500  | 
assumes f: "cont (\<lambda>x. f x)"  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
501  | 
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
 | 
502  | 
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
 | 
503  | 
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
 | 
504  | 
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
 | 
505  | 
|
| 
37079
 
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
 
huffman 
parents: 
36452 
diff
changeset
 | 
506  | 
lemma cont2cont_Let' [simp, cont2cont]:  | 
| 
35933
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
507  | 
assumes f: "cont (\<lambda>x. f x)"  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
508  | 
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
 | 
509  | 
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
 | 
510  | 
using f  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
511  | 
proof (rule cont2cont_Let)  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
512  | 
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
 | 
513  | 
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
 | 
514  | 
next  | 
| 
 
f135ebcc835c
remove continuous let-binding function CLet; add cont2cont rule ordinary Let
 
huffman 
parents: 
35914 
diff
changeset
 | 
515  | 
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
 | 
516  | 
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
 | 
517  | 
qed  | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
17815 
diff
changeset
 | 
518  | 
|
| 39145 | 519  | 
text {* The simple version (suggested by Joachim Breitner) is needed if
 | 
520  | 
the type of the defined term is not a cpo. *}  | 
|
521  | 
||
522  | 
lemma cont2cont_Let_simple [simp, cont2cont]:  | 
|
523  | 
assumes "\<And>y. cont (\<lambda>x. g x y)"  | 
|
524  | 
shows "cont (\<lambda>x. let y = t in g x y)"  | 
|
525  | 
unfolding Let_def using assms .  | 
|
526  | 
||
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
527  | 
end  |