src/HOL/HOLCF/Cont.thy
author wenzelm
Sat, 28 Nov 2020 15:15:53 +0100
changeset 72755 8dffbe01a3e1
parent 68780 54fdc8bc73a3
permissions -rw-r--r--
support for Scala compile-time positions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 40774
diff changeset
     1
(*  Title:      HOL/HOLCF/Cont.thy
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     2
    Author:     Franz Regensburger
35794
8cd7134275cc use headers consistently
huffman
parents: 31902
diff changeset
     3
    Author:     Brian Huffman
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     4
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
     6
section \<open>Continuity and monotonicity\<close>
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     7
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
theory Cont
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
     9
  imports Pcpo
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    11
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    12
text \<open>
15588
14e3228f18cc arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    13
   Now we change the default class! Form now on all untyped type variables are
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2838
diff changeset
    14
   of default class po
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    15
\<close>
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    16
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35914
diff changeset
    17
default_sort po
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    19
subsection \<open>Definitions\<close>
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    20
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67312
diff changeset
    21
definition monofun :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"  \<comment> \<open>monotonicity\<close>
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    22
  where "monofun f \<longleftrightarrow> (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    23
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    24
definition cont :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    25
  where "cont f = (\<forall>Y. chain Y \<longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    26
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    27
lemma contI: "(\<And>Y. chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)) \<Longrightarrow> cont f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    28
  by (simp add: cont_def)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    29
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    30
lemma contE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    31
  by (simp add: cont_def)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    32
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    33
lemma monofunI: "(\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y) \<Longrightarrow> monofun f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    34
  by (simp add: monofun_def)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    35
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    36
lemma monofunE: "monofun f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    37
  by (simp add: monofun_def)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    38
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
    39
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    40
subsection \<open>Equivalence of alternate definition\<close>
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
    41
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    42
text \<open>monotone functions map chains to chains\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    43
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    44
lemma ch2ch_monofun: "monofun f \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. f (Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    45
  apply (rule chainI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    46
  apply (erule monofunE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    47
  apply (erule chainE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    48
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    49
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    50
text \<open>monotone functions map upper bound to upper bounds\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    51
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    52
lemma ub2ub_monofun: "monofun f \<Longrightarrow> range Y <| u \<Longrightarrow> range (\<lambda>i. f (Y i)) <| f u"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    53
  apply (rule ub_rangeI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    54
  apply (erule monofunE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    55
  apply (erule ub_rangeD)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    56
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    57
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    58
text \<open>a lemma about binary chains\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    59
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    60
lemma binchain_cont: "cont f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. f (if i = 0 then x else y)) <<| f y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    61
  apply (subgoal_tac "f (\<Squnion>i::nat. if i = 0 then x else y) = f y")
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    62
   apply (erule subst)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    63
   apply (erule contE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    64
   apply (erule bin_chain)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    65
  apply (rule_tac f=f in arg_cong)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    66
  apply (erule is_lub_bin_chain [THEN lub_eqI])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    67
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    68
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    69
text \<open>continuity implies monotonicity\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    70
16204
5dd79d3f0105 renamed theorems monofun, contlub, cont to monofun_def, etc.; changed intro/elim rules for these predicates into more useful rule_format; removed all MF2 lemmas (Pcpo.thy has more general versions now); cleaned up many proofs.
huffman
parents: 16096
diff changeset
    71
lemma cont2mono: "cont f \<Longrightarrow> monofun f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    72
  apply (rule monofunI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    73
  apply (drule (1) binchain_cont)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    74
  apply (drule_tac i=0 in is_lub_rangeD1)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    75
  apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    76
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    77
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
    78
lemmas cont2monofunE = cont2mono [THEN monofunE]
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
    79
16737
f0fd06dc93e3 add lemmas ch2ch_cont and cont2contlubE
huffman
parents: 16624
diff changeset
    80
lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun]
f0fd06dc93e3 add lemmas ch2ch_cont and cont2contlubE
huffman
parents: 16624
diff changeset
    81
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    82
text \<open>continuity implies preservation of lubs\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    83
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    84
lemma cont2contlubE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    85
  apply (rule lub_eqI [symmetric])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    86
  apply (erule (1) contE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    87
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
    88
25896
b2d2f1ae5808 add lemma contI2
huffman
parents: 25825
diff changeset
    89
lemma contI2:
40736
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    90
  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo"
25896
b2d2f1ae5808 add lemma contI2
huffman
parents: 25825
diff changeset
    91
  assumes mono: "monofun f"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
    92
  assumes below: "\<And>Y. \<lbrakk>chain Y; chain (\<lambda>i. f (Y i))\<rbrakk> \<Longrightarrow> f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))"
25896
b2d2f1ae5808 add lemma contI2
huffman
parents: 25825
diff changeset
    93
  shows "cont f"
40736
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    94
proof (rule contI)
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    95
  fix Y :: "nat \<Rightarrow> 'a"
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    96
  assume Y: "chain Y"
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    97
  with mono have fY: "chain (\<lambda>i. f (Y i))"
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    98
    by (rule ch2ch_monofun)
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
    99
  have "(\<Squnion>i. f (Y i)) = f (\<Squnion>i. Y i)"
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   100
    apply (rule below_antisym)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   101
     apply (rule lub_below [OF fY])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   102
     apply (rule monofunE [OF mono])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   103
     apply (rule is_ub_thelub [OF Y])
40736
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   104
    apply (rule below [OF Y fY])
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   105
    done
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   106
  with fY show "range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)"
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   107
    by (rule thelubE)
72857de90621 isar-style proof for lemma contI2
huffman
parents: 40010
diff changeset
   108
qed
25896
b2d2f1ae5808 add lemma contI2
huffman
parents: 25825
diff changeset
   109
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   110
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   111
subsection \<open>Collection of continuity rules\<close>
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   112
57945
cacb00a569e0 prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
wenzelm
parents: 45294
diff changeset
   113
named_theorems cont2cont "continuity intro rule"
29530
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   114
9905b660612b change to simpler, more extensible continuity simproc
huffman
parents: 29138
diff changeset
   115
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   116
subsection \<open>Continuity of basic functions\<close>
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   117
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   118
text \<open>The identity function is continuous\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   119
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36658
diff changeset
   120
lemma cont_id [simp, cont2cont]: "cont (\<lambda>x. x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   121
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   122
  apply (erule cpo_lubI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   123
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   124
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   125
text \<open>constant functions are continuous\<close>
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   126
37079
0cd15d8c90a0 remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents: 36658
diff changeset
   127
lemma cont_const [simp, cont2cont]: "cont (\<lambda>x. c)"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40736
diff changeset
   128
  using is_lub_const by (rule contI)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   129
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   130
text \<open>application of functions is continuous\<close>
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   131
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   132
lemma cont_apply:
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   133
  fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo" and t :: "'a \<Rightarrow> 'b"
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   134
  assumes 1: "cont (\<lambda>x. t x)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   135
  assumes 2: "\<And>x. cont (\<lambda>y. f x y)"
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   136
  assumes 3: "\<And>y. cont (\<lambda>x. f x y)"
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   137
  shows "cont (\<lambda>x. (f x) (t x))"
35914
91a7311177c4 remove contlub predicate
huffman
parents: 35900
diff changeset
   138
proof (rule contI2 [OF monofunI])
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   139
  fix x y :: "'a"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   140
  assume "x \<sqsubseteq> y"
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   141
  then show "f x (t x) \<sqsubseteq> f y (t y)"
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   142
    by (auto intro: cont2monofunE [OF 1]
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   143
        cont2monofunE [OF 2]
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   144
        cont2monofunE [OF 3]
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   145
        below_trans)
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   146
next
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   147
  fix Y :: "nat \<Rightarrow> 'a"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   148
  assume "chain Y"
35914
91a7311177c4 remove contlub predicate
huffman
parents: 35900
diff changeset
   149
  then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) \<sqsubseteq> (\<Squnion>i. f (Y i) (t (Y i)))"
31041
85b4843d9939 replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents: 31030
diff changeset
   150
    by (simp only: cont2contlubE [OF 1] ch2ch_cont [OF 1]
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   151
        cont2contlubE [OF 2] ch2ch_cont [OF 2]
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   152
        cont2contlubE [OF 3] ch2ch_cont [OF 3]
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   153
        diag_lub below_refl)
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   154
qed
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   155
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   156
lemma cont_compose: "cont c \<Longrightarrow> cont (\<lambda>x. f x) \<Longrightarrow> cont (\<lambda>x. c (f x))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   157
  by (rule cont_apply [OF _ _ cont_const])
29532
59bee7985149 add lemmas cont2monofunE, cont2cont_apply
huffman
parents: 29530
diff changeset
   158
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   159
text \<open>Least upper bounds preserve continuity\<close>
40004
9f6ed6840e8d reformulate lemma cont2cont_lub and move to Cont.thy
huffman
parents: 37099
diff changeset
   160
9f6ed6840e8d reformulate lemma cont2cont_lub and move to Cont.thy
huffman
parents: 37099
diff changeset
   161
lemma cont2cont_lub [simp]:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   162
  assumes chain: "\<And>x. chain (\<lambda>i. F i x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   163
    and cont: "\<And>i. cont (\<lambda>x. F i x)"
40004
9f6ed6840e8d reformulate lemma cont2cont_lub and move to Cont.thy
huffman
parents: 37099
diff changeset
   164
  shows "cont (\<lambda>x. \<Squnion>i. F i x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   165
  apply (rule contI2)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   166
   apply (simp add: monofunI cont2monofunE [OF cont] lub_mono chain)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   167
  apply (simp add: cont2contlubE [OF cont])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   168
  apply (simp add: diag_lub ch2ch_cont [OF cont] chain)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   169
  done
40004
9f6ed6840e8d reformulate lemma cont2cont_lub and move to Cont.thy
huffman
parents: 37099
diff changeset
   170
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   171
text \<open>if-then-else is continuous\<close>
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   172
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   173
lemma cont_if [simp, cont2cont]: "cont f \<Longrightarrow> cont g \<Longrightarrow> cont (\<lambda>x. if b then f x else g x)"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   174
  by (induct b) simp_all
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   175
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   176
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   177
subsection \<open>Finite chains and flat pcpos\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   178
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   179
text \<open>Monotone functions map finite chains to finite chains.\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   180
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   181
lemma monofun_finch2finch: "monofun f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   182
  by (force simp add: finite_chain_def ch2ch_monofun max_in_chain_def)
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   183
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   184
text \<open>The same holds for continuous functions.\<close>
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   185
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   186
lemma cont_finch2finch: "cont f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   187
  by (rule cont2mono [THEN monofun_finch2finch])
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   188
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   189
text \<open>All monotone functions with chain-finite domain are continuous.\<close>
40010
d7fdd84b959f edit comments
huffman
parents: 40004
diff changeset
   190
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   191
lemma chfindom_monofun2cont: "monofun f \<Longrightarrow> cont f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   192
  for f :: "'a::chfin \<Rightarrow> 'b::cpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   193
  apply (erule contI2)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   194
  apply (frule chfin2finch)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   195
  apply (clarsimp simp add: finite_chain_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   196
  apply (subgoal_tac "max_in_chain i (\<lambda>i. f (Y i))")
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   197
   apply (simp add: maxinch_is_thelub ch2ch_monofun)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   198
  apply (force simp add: max_in_chain_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   199
  done
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   200
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   201
text \<open>All strict functions with flat domain are continuous.\<close>
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   202
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   203
lemma flatdom_strict2mono: "f \<bottom> = \<bottom> \<Longrightarrow> monofun f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   204
  for f :: "'a::flat \<Rightarrow> 'b::pcpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   205
  apply (rule monofunI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   206
  apply (drule ax_flat)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   207
  apply auto
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   208
  done
16624
645b9560f3fd cleaned up; reorganized and added section headings
huffman
parents: 16564
diff changeset
   209
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   210
lemma flatdom_strict2cont: "f \<bottom> = \<bottom> \<Longrightarrow> cont f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   211
  for f :: "'a::flat \<Rightarrow> 'b::pcpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   212
  by (rule flatdom_strict2mono [THEN chfindom_monofun2cont])
15565
2454493bd77b converted to new-style theory
huffman
parents: 14981
diff changeset
   213
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
   214
text \<open>All functions with discrete domain are continuous.\<close>
26024
d5129e687290 new lemma cont_discrete_cpo
huffman
parents: 25920
diff changeset
   215
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   216
lemma cont_discrete_cpo [simp, cont2cont]: "cont f"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   217
  for f :: "'a::discrete_cpo \<Rightarrow> 'b::cpo"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   218
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   219
  apply (drule discrete_chain_const, clarify)
68780
54fdc8bc73a3 new simp rule
haftmann
parents: 67443
diff changeset
   220
  apply simp
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 62175
diff changeset
   221
  done
26024
d5129e687290 new lemma cont_discrete_cpo
huffman
parents: 25920
diff changeset
   222
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   223
end