src/HOLCF/Fix.thy
author huffman
Thu, 03 Nov 2005 00:43:11 +0100
changeset 18074 a92b7c5133de
parent 17816 9942c5ed866a
child 18078 20e5a6440790
permissions -rw-r--r--
reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2275
diff changeset
     1
(*  Title:      HOLCF/Fix.thy
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1410
diff changeset
     3
    Author:     Franz Regensburger
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     4
16070
4a83dd540b88 removed LICENCE note -- everything is subject to Isabelle licence as
wenzelm
parents: 16056
diff changeset
     5
Definitions for fixed point operator and admissibility.
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     6
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
header {* Fixed point operator and admissibility *}
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
theory Fix
16056
32c3b7188c28 Moved admissibility definitions and lemmas to a separate theory
huffman
parents: 16006
diff changeset
    11
imports Cfun Cprod Adm
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    12
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    13
16082
ebb53ebfd4e2 added defaultsort declaration
huffman
parents: 16079
diff changeset
    14
defaultsort pcpo
ebb53ebfd4e2 added defaultsort declaration
huffman
parents: 16079
diff changeset
    15
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    16
subsection {* Definitions *}
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    17
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
consts
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    19
  iterate :: "nat \<Rightarrow> ('a \<rightarrow> 'a) \<rightarrow> 'a \<rightarrow> 'a"
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
    20
  "fix"   :: "('a \<rightarrow> 'a) \<rightarrow> 'a"
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
    21
  admw    :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    22
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4721
diff changeset
    23
primrec
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    24
  "iterate 0 = (\<Lambda> F x. x)"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    25
  "iterate (Suc n) = (\<Lambda> F x. F\<cdot>(iterate n\<cdot>F\<cdot>x))"
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4721
diff changeset
    26
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    27
defs
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    28
  fix_def:       "fix \<equiv> \<Lambda> F. \<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    29
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    30
  admw_def:      "admw P \<equiv> \<forall>F. (\<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)) \<longrightarrow>
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    31
                            P (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>)" 
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    32
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    33
subsection {* Binder syntax for @{term fix} *}
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    34
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    35
syntax
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17585
diff changeset
    36
  "_FIX" :: "[pttrn, 'a] => 'a" ("(3FIX _./ _)" [0, 10] 10)
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    37
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    38
syntax (xsymbols)
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17585
diff changeset
    39
  "_FIX" :: "[pttrn, 'a] => 'a" ("(3\<mu>_./ _)" [0, 10] 10)
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    40
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    41
translations
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 17585
diff changeset
    42
  "FIX x. t" == "fix$(LAM x. t)"
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    43
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    44
subsection {* Properties of @{term iterate} *}
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    45
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    46
text {* derive inductive properties of iterate from primitive recursion *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    47
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    48
lemma iterate_0 [simp]: "iterate 0\<cdot>F\<cdot>x = x"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    49
by simp
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    50
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    51
lemma iterate_Suc [simp]: "iterate (Suc n)\<cdot>F\<cdot>x = F\<cdot>(iterate n\<cdot>F\<cdot>x)"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    52
by simp
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    53
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    54
declare iterate.simps [simp del]
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    55
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    56
lemma iterate_Suc2: "iterate (Suc n)\<cdot>F\<cdot>x = iterate n\<cdot>F\<cdot>(F\<cdot>x)"
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
    57
by (induct_tac n, auto)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    58
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    59
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    60
  The sequence of function iterations is a chain.
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    61
  This property is essential since monotonicity of iterate makes no sense.
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    62
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    63
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    64
lemma chain_iterate2: "x \<sqsubseteq> F\<cdot>x \<Longrightarrow> chain (\<lambda>i. iterate i\<cdot>F\<cdot>x)"
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
    65
by (rule chainI, induct_tac i, auto elim: monofun_cfun_arg)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    66
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    67
lemma chain_iterate [simp]: "chain (\<lambda>i. iterate i\<cdot>F\<cdot>\<bottom>)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    68
by (rule chain_iterate2 [OF minimal])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    69
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    70
subsection {* Properties of @{term fix} *}
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    71
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    72
text {* direct connection between @{term fix} and iteration without @{term Ifix} *}
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    73
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    74
lemma fix_def2: "fix\<cdot>F = (\<Squnion>i. iterate i\<cdot>F\<cdot>\<bottom>)"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    75
apply (unfold fix_def)
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    76
apply (rule beta_cfun)
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    77
apply (rule cont2cont_lub)
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    78
apply (rule ch2ch_fun_rev)
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    79
apply (rule chain_iterate)
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    80
apply simp
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    81
done
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    82
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    83
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    84
  Kleene's fixed point theorems for continuous functions in pointed
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    85
  omega cpo's
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    86
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    87
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    88
lemma fix_eq: "fix\<cdot>F = F\<cdot>(fix\<cdot>F)"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    89
apply (simp add: fix_def2)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    90
apply (subst lub_range_shift [of _ 1, symmetric])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    91
apply (rule chain_iterate)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    92
apply (subst contlub_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    93
apply (rule chain_iterate)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    94
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    95
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    96
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    97
lemma fix_least_less: "F\<cdot>x \<sqsubseteq> x \<Longrightarrow> fix\<cdot>F \<sqsubseteq> x"
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
    98
apply (simp add: fix_def2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    99
apply (rule is_lub_thelub)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   100
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   101
apply (rule ub_rangeI)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   102
apply (induct_tac i)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   103
apply simp
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   104
apply simp
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
   105
apply (erule rev_trans_less)
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   106
apply (erule monofun_cfun_arg)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   107
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   108
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   109
lemma fix_least: "F\<cdot>x = x \<Longrightarrow> fix\<cdot>F \<sqsubseteq> x"
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
   110
by (rule fix_least_less, simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   111
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   112
lemma fix_eqI: "\<lbrakk>F\<cdot>x = x; \<forall>z. F\<cdot>z = z \<longrightarrow> x \<sqsubseteq> z\<rbrakk> \<Longrightarrow> x = fix\<cdot>F"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   113
apply (rule antisym_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   114
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   115
apply (erule mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   116
apply (rule fix_eq [symmetric])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   117
apply (erule fix_least)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   118
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   119
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   120
lemma fix_eq2: "f \<equiv> fix\<cdot>F \<Longrightarrow> f = F\<cdot>f"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   121
by (simp add: fix_eq [symmetric])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   122
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   123
lemma fix_eq3: "f \<equiv> fix\<cdot>F \<Longrightarrow> f\<cdot>x = F\<cdot>f\<cdot>x"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   124
by (erule fix_eq2 [THEN cfun_fun_cong])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   125
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   126
lemma fix_eq4: "f = fix\<cdot>F \<Longrightarrow> f = F\<cdot>f"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   127
apply (erule ssubst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   128
apply (rule fix_eq)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   129
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   130
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   131
lemma fix_eq5: "f = fix\<cdot>F \<Longrightarrow> f\<cdot>x = F\<cdot>f\<cdot>x"
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   132
by (erule fix_eq4 [THEN cfun_fun_cong])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   133
16556
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   134
text {* strictness of @{term fix} *}
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   135
16917
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   136
lemma fix_defined_iff: "(fix\<cdot>F = \<bottom>) = (F\<cdot>\<bottom> = \<bottom>)"
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   137
apply (rule iffI)
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   138
apply (erule subst)
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   139
apply (rule fix_eq [symmetric])
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   140
apply (erule fix_least [THEN UU_I])
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   141
done
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   142
16556
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   143
lemma fix_strict: "F\<cdot>\<bottom> = \<bottom> \<Longrightarrow> fix\<cdot>F = \<bottom>"
16917
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   144
by (simp add: fix_defined_iff)
16556
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   145
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   146
lemma fix_defined: "F\<cdot>\<bottom> \<noteq> \<bottom> \<Longrightarrow> fix\<cdot>F \<noteq> \<bottom>"
16917
1fe50b19daba add theorem fix_defined_iff; cleaned up
huffman
parents: 16556
diff changeset
   147
by (simp add: fix_defined_iff)
16556
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   148
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   149
text {* @{term fix} applied to identity and constant functions *}
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   150
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   151
lemma fix_id: "(\<mu> x. x) = \<bottom>"
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   152
by (simp add: fix_strict)
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   153
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   154
lemma fix_const: "(\<mu> x. c) = c"
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
   155
by (subst fix_eq, simp)
16556
a0c8d0499b5f added theorems fix_strict, fix_defined, fix_id, fix_const
huffman
parents: 16388
diff changeset
   156
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   157
subsection {* Admissibility and fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   158
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   159
text {* an admissible formula is also weak admissible *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   160
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   161
lemma adm_impl_admw: "adm P \<Longrightarrow> admw P"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   162
apply (unfold admw_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   163
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   164
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   165
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   166
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   167
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   168
16079
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   169
text {* some lemmata for functions with flat/chfin domain/range types *}
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   170
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   171
lemma adm_chfindom: "adm (\<lambda>(u::'a::cpo \<rightarrow> 'b::chfin). P(u\<cdot>s))"
16079
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   172
apply (unfold adm_def)
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   173
apply (intro strip)
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   174
apply (drule chfin_Rep_CFunR)
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   175
apply (erule_tac x = "s" in allE)
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   176
apply clarsimp
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   177
done
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   178
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   179
(* adm_flat not needed any more, since it is a special case of adm_chfindom *)
757e1c4a8081 moved adm_chfindom from Adm.thy to Fix.thy, to remove dependence on Cfun
huffman
parents: 16070
diff changeset
   180
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   181
text {* fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   182
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   183
lemma fix_ind: "\<lbrakk>adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F\<cdot>x)\<rbrakk> \<Longrightarrow> P (fix\<cdot>F)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   184
apply (subst fix_def2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   185
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   186
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   187
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   188
apply (induct_tac "i")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   189
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   190
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   191
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   192
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   193
lemma def_fix_ind:
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   194
  "\<lbrakk>f \<equiv> fix\<cdot>F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F\<cdot>x)\<rbrakk> \<Longrightarrow> P f"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   195
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   196
apply (erule fix_ind)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   197
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   198
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   199
done
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   200
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   201
text {* computational induction for weak admissible formulae *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   202
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
   203
lemma wfix_ind: "\<lbrakk>admw P; \<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)\<rbrakk> \<Longrightarrow> P (fix\<cdot>F)"
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   204
by (simp add: fix_def2 admw_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   205
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   206
lemma def_wfix_ind:
18074
a92b7c5133de reorganized; removed intermediate constant Ifix; changed iterate to a continuous type; added theorem fix_least_less
huffman
parents: 17816
diff changeset
   207
  "\<lbrakk>f \<equiv> fix\<cdot>F; admw P; \<forall>n. P (iterate n\<cdot>F\<cdot>\<bottom>)\<rbrakk> \<Longrightarrow> P f"
16214
e3816a7db016 cleaned up proof of cont_Ifix
huffman
parents: 16082
diff changeset
   208
by (simp, rule wfix_ind)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   209
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   210
end