src/HOLCF/Fix.thy
author wenzelm
Wed, 25 May 2005 09:44:34 +0200
changeset 16070 4a83dd540b88
parent 16056 32c3b7188c28
child 16079 757e1c4a8081
permissions -rw-r--r--
removed LICENCE note -- everything is subject to Isabelle licence as stated in COPYRIGHT file;
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
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    14
subsection {* Definitions *}
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    15
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    16
consts
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    17
1990
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    18
iterate	:: "nat=>('a->'a)=>'a=>'a"
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    19
Ifix	:: "('a->'a)=>'a"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    20
"fix"	:: "('a->'a)->'a"
1990
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    21
admw		:: "('a=>bool)=>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
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    24
  iterate_0:   "iterate 0 F x = x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    25
  iterate_Suc: "iterate (Suc n) F x  = F$(iterate n F 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
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    28
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    29
Ifix_def:      "Ifix F == lub(range(%i. iterate i F UU))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    30
fix_def:       "fix == (LAM f. Ifix f)"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    31
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    32
admw_def:      "admw P == !F. (!n. P (iterate n F UU)) -->
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1410
diff changeset
    33
                            P (lub(range (%i. iterate i F UU)))" 
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    34
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    35
subsection {* Binder syntax for @{term fix} *}
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    36
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    37
syntax
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    38
  "@FIX" :: "('a => 'a) => 'a"  (binder "FIX " 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    39
  "@FIXP" :: "[patterns, 'a] => 'a"  ("(3FIX <_>./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    40
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    41
syntax (xsymbols)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    42
  "FIX " :: "[idt, 'a] => 'a"  ("(3\<mu>_./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    43
  "@FIXP" :: "[patterns, 'a] => 'a"  ("(3\<mu>()<_>./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    44
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    45
translations
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    46
  "FIX x. LAM y. t" == "fix\<cdot>(LAM x y. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    47
  "FIX x. t" == "fix\<cdot>(LAM x. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    48
  "FIX <xs>. t" == "fix\<cdot>(LAM <xs>. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    49
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    50
subsection {* Properties of @{term iterate} and @{term fix} *}
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    51
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    52
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
    53
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    54
lemma iterate_Suc2: "iterate (Suc n) F x = iterate n F (F$x)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    55
by (induct_tac "n", auto)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    56
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    57
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    58
  The sequence of function iterations is a chain.
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    59
  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
    60
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    61
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    62
lemma chain_iterate2: "x << F$x ==> chain (%i. iterate i F x)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    63
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
    64
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    65
lemma chain_iterate: "chain (%i. iterate i F UU)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    66
by (rule chain_iterate2 [OF minimal])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    67
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    68
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    69
  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
    70
  omega cpo's
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    71
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    72
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    73
lemma Ifix_eq: "Ifix F = F$(Ifix F)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    74
apply (unfold Ifix_def)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    75
apply (subst lub_range_shift [of _ 1, symmetric])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    76
apply (rule chain_iterate)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    77
apply (subst contlub_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    78
apply (rule chain_iterate)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    79
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    80
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    81
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    82
lemma Ifix_least: "F$x=x ==> Ifix(F) << x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    83
apply (unfold Ifix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    84
apply (rule is_lub_thelub)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    85
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    86
apply (rule ub_rangeI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    87
apply (induct_tac "i")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    88
apply (simp (no_asm_simp))
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    89
apply (simp (no_asm_simp))
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    90
apply (erule_tac t = "x" in subst)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    91
apply (erule monofun_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    92
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    93
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    94
text {* monotonicity and continuity of @{term iterate} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    95
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    96
lemma cont_iterate: "cont(iterate(i))"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    97
apply (induct_tac i)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    98
apply simp
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    99
apply simp
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   100
apply (rule cont2cont_CF1L_rev)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   101
apply (rule allI)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   102
apply (rule cont2cont_Rep_CFun)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   103
apply (rule cont_id)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   104
apply (erule cont2cont_CF1L)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   105
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   106
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   107
lemma monofun_iterate: "monofun(iterate(i))"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   108
by (rule cont_iterate [THEN cont2mono])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   109
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   110
lemma contlub_iterate: "contlub(iterate(i))"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   111
by (rule cont_iterate [THEN cont2contlub])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   112
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   113
text {* a lemma about continuity of @{term iterate} in its third argument *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   114
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   115
lemma cont_iterate2: "cont (iterate n F)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   116
by (induct_tac "n", simp_all)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   117
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   118
lemma monofun_iterate2: "monofun(iterate n F)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   119
by (rule cont_iterate2 [THEN cont2mono])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   120
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   121
lemma contlub_iterate2: "contlub(iterate n F)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   122
by (rule cont_iterate2 [THEN cont2contlub])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   123
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   124
text {* monotonicity and continuity of @{term Ifix} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   125
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   126
text {* better access to definitions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   127
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   128
lemma Ifix_def2: "Ifix=(%x. lub(range(%i. iterate i x UU)))"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   129
apply (rule ext)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   130
apply (unfold Ifix_def)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   131
apply (rule refl)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   132
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   133
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   134
lemma cont_Ifix: "cont(Ifix)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   135
apply (subst Ifix_def2)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   136
apply (subst cont_iterate [THEN cont2cont_CF1L, THEN beta_cfun, symmetric])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   137
apply (rule cont_lubcfun)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   138
apply (rule chainI)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   139
apply (rule less_cfun2)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   140
apply (simp add: cont_iterate [THEN cont2cont_CF1L] del: iterate_Suc)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   141
apply (rule chainE)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   142
apply (rule chain_iterate)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   143
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   144
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   145
lemma monofun_Ifix: "monofun(Ifix)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   146
by (rule cont_Ifix [THEN cont2mono])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   147
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   148
lemma contlub_Ifix: "contlub(Ifix)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   149
by (rule cont_Ifix [THEN cont2contlub])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   150
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   151
text {* propagate properties of @{term Ifix} to its continuous counterpart *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   152
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   153
lemma fix_eq: "fix$F = F$(fix$F)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   154
apply (unfold fix_def)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   155
apply (simp add: cont_Ifix)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   156
apply (rule Ifix_eq)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   157
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   158
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   159
lemma fix_least: "F$x = x ==> fix$F << x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   160
apply (unfold fix_def)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   161
apply (simp add: cont_Ifix)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   162
apply (erule Ifix_least)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   163
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   164
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   165
lemma fix_eqI: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   166
"[| F$x = x; !z. F$z = z --> x << z |] ==> x = fix$F"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   167
apply (rule antisym_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   168
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   169
apply (erule mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   170
apply (rule fix_eq [symmetric])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   171
apply (erule fix_least)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   172
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   173
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   174
lemma fix_eq2: "f == fix$F ==> f = F$f"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   175
by (simp add: fix_eq [symmetric])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   176
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   177
lemma fix_eq3: "f == fix$F ==> f$x = F$f$x"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   178
by (erule fix_eq2 [THEN cfun_fun_cong])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   179
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   180
(* fun fix_tac3 thm i  = ((rtac trans i) THEN (rtac (thm RS fix_eq3) i)) *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   181
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   182
lemma fix_eq4: "f = fix$F ==> f = F$f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   183
apply (erule ssubst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   184
apply (rule fix_eq)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   185
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   186
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   187
lemma fix_eq5: "f = fix$F ==> f$x = F$f$x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   188
apply (rule trans)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   189
apply (erule fix_eq4 [THEN cfun_fun_cong])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   190
apply (rule refl)
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
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   193
(* fun fix_tac5 thm i  = ((rtac trans i) THEN (rtac (thm RS fix_eq5) i)) *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   194
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   195
(* proves the unfolding theorem for function equations f = fix$... *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   196
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   197
fun fix_prover thy fixeq s = prove_goal thy s (fn prems => [
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   198
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   199
        (rtac (fixeq RS fix_eq4) 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   200
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   201
        (rtac beta_cfun 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   202
        (Simp_tac 1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   203
        ])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   204
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   205
(* proves the unfolding theorem for function definitions f == fix$... *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   206
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   207
fun fix_prover2 thy fixdef s = prove_goal thy s (fn prems => [
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   208
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   209
        (rtac (fix_eq2) 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   210
        (rtac fixdef 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   211
        (rtac beta_cfun 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   212
        (Simp_tac 1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   213
        ])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   214
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   215
(* proves an application case for a function from its unfolding thm *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   216
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   217
fun case_prover thy unfold s = prove_goal thy s (fn prems => [
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   218
	(cut_facts_tac prems 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   219
	(rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   220
	(stac unfold 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   221
	Auto_tac
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   222
	])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   223
*)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   224
text {* direct connection between @{term fix} and iteration without @{term Ifix} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   225
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   226
lemma fix_def2: "fix$F = lub(range(%i. iterate i F UU))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   227
apply (unfold fix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   228
apply (fold Ifix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   229
apply (simp (no_asm_simp) add: cont_Ifix)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   230
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   231
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   232
subsection {* Admissibility and fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   233
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   234
lemma admw_def2: "admw(P) = (!F.(!n. P(iterate n F UU)) --> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   235
                         P (lub(range(%i. iterate i F UU))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   236
apply (unfold admw_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   237
apply (rule refl)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   238
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   239
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   240
text {* an admissible formula is also weak admissible *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   241
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   242
lemma adm_impl_admw: "adm(P)==>admw(P)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   243
apply (unfold admw_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   244
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   245
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   246
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   247
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   248
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   249
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   250
text {* fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   251
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   252
lemma fix_ind:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   253
     "[| adm(P); P(UU); !!x. P(x) ==> P(F$x)|] ==> P(fix$F)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   254
apply (subst fix_def2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   255
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   256
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   257
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   258
apply (induct_tac "i")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   259
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   260
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   261
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   262
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   263
lemma def_fix_ind: "[| f == fix$F; adm(P);  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   264
        P(UU); !!x. P(x) ==> P(F$x)|] ==> P f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   265
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   266
apply (erule fix_ind)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   267
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   268
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   269
done
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   270
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   271
text {* computational induction for weak admissible formulae *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   272
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   273
lemma wfix_ind: "[| admw(P); !n. P(iterate n F UU)|] ==> P(fix$F)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   274
apply (subst fix_def2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   275
apply (rule admw_def2 [THEN iffD1, THEN spec, THEN mp])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   276
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   277
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   278
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   279
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   280
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   281
lemma def_wfix_ind: "[| f == fix$F; admw(P);  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   282
        !n. P(iterate n F UU) |] ==> P f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   283
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   284
apply (erule wfix_ind)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   285
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   286
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   287
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   288
end
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   289