src/HOLCF/Fix.thy
author huffman
Wed, 18 May 2005 23:29:36 +0200
changeset 16005 42f3f299ee68
parent 15857 83cb9dae3817
child 16006 693dd363e0bf
permissions -rw-r--r--
cleaned up and shortened some proofs
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
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     6
definitions for fixed point operator and admissibility
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     8
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
header {* Fixed point operator and admissibility *}
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    11
theory Fix
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    12
imports Cfun Cprod
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    13
begin
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    14
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    15
subsection {* Definitions *}
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    16
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    17
consts
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
1990
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    19
iterate	:: "nat=>('a->'a)=>'a=>'a"
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    20
Ifix	:: "('a->'a)=>'a"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    21
"fix"	:: "('a->'a)->'a"
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2640
diff changeset
    22
adm		:: "('a::cpo=>bool)=>bool"
1990
9e23119c0219 added comment on is_flat
oheimb
parents: 1825
diff changeset
    23
admw		:: "('a=>bool)=>bool"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    24
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4721
diff changeset
    25
primrec
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    26
  iterate_0:   "iterate 0 F x = x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    27
  iterate_Suc: "iterate (Suc n) F x  = F$(iterate n F x)"
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4721
diff changeset
    28
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    29
defs
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    30
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    31
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
    32
fix_def:       "fix == (LAM f. Ifix f)"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    33
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    34
adm_def:       "adm P == !Y. chain(Y) --> 
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3326
diff changeset
    35
                        (!i. P(Y i)) --> P(lub(range Y))"
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    36
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    37
admw_def:      "admw P == !F. (!n. P (iterate n F UU)) -->
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1410
diff changeset
    38
                            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
    39
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    40
subsection {* Binder syntax for @{term fix} *}
15857
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    41
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    42
syntax
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    43
  "@FIX" :: "('a => 'a) => 'a"  (binder "FIX " 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    44
  "@FIXP" :: "[patterns, 'a] => 'a"  ("(3FIX <_>./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    45
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    46
syntax (xsymbols)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    47
  "FIX " :: "[idt, 'a] => 'a"  ("(3\<mu>_./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    48
  "@FIXP" :: "[patterns, 'a] => 'a"  ("(3\<mu>()<_>./ _)" [0, 10] 10)
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    49
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    50
translations
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    51
  "FIX x. LAM y. t" == "fix\<cdot>(LAM x y. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    52
  "FIX x. t" == "fix\<cdot>(LAM x. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    53
  "FIX <xs>. t" == "fix\<cdot>(LAM <xs>. t)"
83cb9dae3817 Added binder syntax for fix
huffman
parents: 15637
diff changeset
    54
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    55
subsection {* Properties of @{term iterate} and @{term fix} *}
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    56
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    57
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
    58
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    59
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
    60
by (induct_tac "n", auto)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    61
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    62
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    63
  The sequence of function iterations is a chain.
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    64
  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
    65
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    66
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    67
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
    68
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
    69
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    70
lemma chain_iterate: "chain (%i. iterate i F UU)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    71
by (rule chain_iterate2 [OF minimal])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    72
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    73
text {*
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    74
  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
    75
  omega cpo's
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    76
*}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    77
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    78
lemma Ifix_eq: "Ifix F = F$(Ifix F)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    79
apply (unfold Ifix_def)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    80
apply (subst lub_range_shift [of _ 1, symmetric])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    81
apply (rule chain_iterate)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    82
apply (subst contlub_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    83
apply (rule chain_iterate)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    84
apply simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    85
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    86
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    87
lemma Ifix_least: "F$x=x ==> Ifix(F) << x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    88
apply (unfold Ifix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    89
apply (rule is_lub_thelub)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    90
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    91
apply (rule ub_rangeI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    92
apply (induct_tac "i")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    93
apply (simp (no_asm_simp))
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    94
apply (simp (no_asm_simp))
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
    95
apply (erule_tac t = "x" in subst)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    96
apply (erule monofun_cfun_arg)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    97
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
    98
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
    99
text {* monotonicity and continuity of @{term iterate} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   100
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   101
lemma cont_iterate: "cont(iterate(i))"
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
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   105
apply (rule cont2cont_CF1L_rev)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   106
apply (rule allI)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   107
apply (rule cont2cont_Rep_CFun)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   108
apply (rule cont_id)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   109
apply (erule cont2cont_CF1L)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   110
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   111
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   112
lemma monofun_iterate: "monofun(iterate(i))"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   113
by (rule cont_iterate [THEN cont2mono])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   114
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   115
lemma contlub_iterate: "contlub(iterate(i))"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   116
by (rule cont_iterate [THEN cont2contlub])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   117
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   118
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
   119
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   120
lemma cont_iterate2: "cont (iterate n F)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   121
by (induct_tac "n", simp_all)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   122
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   123
lemma monofun_iterate2: "monofun(iterate n F)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   124
by (rule cont_iterate2 [THEN cont2mono])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   125
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   126
lemma contlub_iterate2: "contlub(iterate n F)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   127
by (rule cont_iterate2 [THEN cont2contlub])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   128
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   129
text {* monotonicity and continuity of @{term Ifix} *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   130
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   131
text {* better access to definitions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   132
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   133
lemma Ifix_def2: "Ifix=(%x. lub(range(%i. iterate i x UU)))"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   134
apply (rule ext)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   135
apply (unfold Ifix_def)
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   136
apply (rule refl)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   137
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   138
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   139
lemma cont_Ifix: "cont(Ifix)"
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   140
apply (subst Ifix_def2)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   141
apply (subst cont_iterate [THEN cont2cont_CF1L, THEN beta_cfun, symmetric])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   142
apply (rule cont_lubcfun)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   143
apply (rule chainI)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   144
apply (rule less_cfun2)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   145
apply (simp add: cont_iterate [THEN cont2cont_CF1L] del: iterate_Suc)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   146
apply (rule chainE)
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   147
apply (rule chain_iterate)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   148
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   149
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   150
lemma monofun_Ifix: "monofun(Ifix)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   151
by (rule cont_Ifix [THEN cont2mono])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   152
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   153
lemma contlub_Ifix: "contlub(Ifix)"
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   154
by (rule cont_Ifix [THEN cont2contlub])
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   155
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   156
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
   157
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   158
lemma fix_eq: "fix$F = F$(fix$F)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   159
apply (unfold fix_def)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   160
apply (simp add: cont_Ifix)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   161
apply (rule Ifix_eq)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   162
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   163
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   164
lemma fix_least: "F$x = x ==> fix$F << x"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   165
apply (unfold fix_def)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   166
apply (simp add: cont_Ifix)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   167
apply (erule Ifix_least)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   168
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   169
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   170
lemma fix_eqI: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   171
"[| 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
   172
apply (rule antisym_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   173
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   174
apply (erule mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   175
apply (rule fix_eq [symmetric])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   176
apply (erule fix_least)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   177
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   178
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   179
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
   180
by (simp add: fix_eq [symmetric])
15576
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_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
   183
by (erule fix_eq2 [THEN cfun_fun_cong])
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   184
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   185
(* 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
   186
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   187
lemma fix_eq4: "f = fix$F ==> f = F$f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   188
apply (erule ssubst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   189
apply (rule fix_eq)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   190
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   191
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   192
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
   193
apply (rule trans)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   194
apply (erule fix_eq4 [THEN cfun_fun_cong])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   195
apply (rule refl)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   196
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   197
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   198
(* 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
   199
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   200
(* proves the unfolding theorem for function equations f = fix$... *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   201
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   202
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
   203
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   204
        (rtac (fixeq RS fix_eq4) 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   205
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   206
        (rtac beta_cfun 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   207
        (Simp_tac 1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   208
        ])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   209
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   210
(* proves the unfolding theorem for function definitions f == fix$... *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   211
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   212
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
   213
        (rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   214
        (rtac (fix_eq2) 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   215
        (rtac fixdef 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   216
        (rtac beta_cfun 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   217
        (Simp_tac 1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   218
        ])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   219
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   220
(* 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
   221
(*
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   222
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
   223
	(cut_facts_tac prems 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   224
	(rtac trans 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   225
	(stac unfold 1),
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   226
	Auto_tac
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   227
	])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   228
*)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   229
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
   230
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   231
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
   232
apply (unfold fix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   233
apply (fold Ifix_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   234
apply (simp (no_asm_simp) add: cont_Ifix)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   235
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   236
16005
42f3f299ee68 cleaned up and shortened some proofs
huffman
parents: 15857
diff changeset
   237
subsection {* Admissibility and fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   238
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   239
text {* access to definitions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   240
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   241
lemma admI:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   242
   "(!!Y. [| chain Y; !i. P (Y i) |] ==> P (lub (range Y))) ==> adm P"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   243
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   244
apply blast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   245
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   246
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   247
lemma triv_admI: "!x. P x ==> adm P"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   248
apply (rule admI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   249
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   250
done
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 admD: "[| adm(P); chain(Y); !i. P(Y(i)) |] ==> P(lub(range(Y)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   253
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   254
apply blast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   255
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   256
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   257
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
   258
                         P (lub(range(%i. iterate i F UU))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   259
apply (unfold admw_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   260
apply (rule refl)
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
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   263
text {* an admissible formula is also weak admissible *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   264
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   265
lemma adm_impl_admw: "adm(P)==>admw(P)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   266
apply (unfold admw_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   267
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   268
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   269
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   270
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   271
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   272
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   273
text {* fixed point induction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   274
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   275
lemma fix_ind:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   276
     "[| 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
   277
apply (subst fix_def2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   278
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   279
apply (rule chain_iterate)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   280
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   281
apply (induct_tac "i")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   282
apply simp
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
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   285
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   286
lemma def_fix_ind: "[| f == fix$F; adm(P);  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   287
        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
   288
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   289
apply (erule fix_ind)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   290
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   291
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   292
done
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   293
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   294
text {* computational induction for weak admissible formulae *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   295
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   296
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
   297
apply (subst fix_def2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   298
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
   299
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   300
apply (rule allI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   301
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   302
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   303
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   304
lemma def_wfix_ind: "[| f == fix$F; admw(P);  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   305
        !n. P(iterate n F UU) |] ==> P f"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   306
apply simp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   307
apply (erule wfix_ind)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   308
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   309
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   310
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   311
text {* for chain-finite (easy) types every formula is admissible *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   312
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   313
lemma adm_max_in_chain: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   314
"!Y. chain(Y::nat=>'a) --> (? n. max_in_chain n Y) ==> adm(P::'a=>bool)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   315
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   316
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   317
apply (rule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   318
apply (rule mp)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   319
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   320
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   321
apply (subst lub_finch1 [THEN thelubI])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   322
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   323
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   324
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   325
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   326
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   327
lemmas adm_chfin = chfin [THEN adm_max_in_chain, standard]
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   328
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   329
text {* some lemmata for functions with flat/chfin domain/range types *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   330
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   331
lemma adm_chfindom: "adm (%(u::'a::cpo->'b::chfin). P(u$s))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   332
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   333
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   334
apply (drule chfin_Rep_CFunR)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   335
apply (erule_tac x = "s" in allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   336
apply clarsimp
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   337
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   338
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   339
(* adm_flat not needed any more, since it is a special case of adm_chfindom *)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   340
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   341
text {* improved admissibility introduction *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   342
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   343
lemma admI2:
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   344
 "(!!Y. [| chain Y; !i. P (Y i); !i. ? j. i < j & Y i ~= Y j & Y i << Y j |] 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   345
  ==> P(lub (range Y))) ==> adm P"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   346
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   347
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   348
apply (erule increasing_chain_adm_lemma)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   349
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   350
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   351
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   352
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   353
text {* admissibility of special formulae and propagation *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   354
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   355
lemma adm_less [simp]: "[|cont u;cont v|]==> adm(%x. u x << v x)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   356
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   357
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   358
apply (frule_tac f = "u" in cont2mono [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   359
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   360
apply (frule_tac f = "v" in cont2mono [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   361
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   362
apply (erule cont2contlub [THEN contlubE, THEN spec, THEN mp, THEN ssubst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   363
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   364
apply (erule cont2contlub [THEN contlubE, THEN spec, THEN mp, THEN ssubst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   365
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   366
apply (blast intro: lub_mono)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   367
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   368
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   369
lemma adm_conj [simp]: "[| adm P; adm Q |] ==> adm(%x. P x & Q x)"
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   370
by (fast elim: admD intro: admI)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   371
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   372
lemma adm_not_free [simp]: "adm(%x. t)"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   373
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   374
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   375
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   376
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   377
lemma adm_not_less: "cont t ==> adm(%x.~ (t x) << u)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   378
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   379
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   380
apply (rule contrapos_nn)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   381
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   382
apply (rule trans_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   383
prefer 2 apply (assumption)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   384
apply (erule cont2mono [THEN monofun_fun_arg])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   385
apply (rule is_ub_thelub)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   386
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   387
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   388
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   389
lemma adm_all: "!y. adm(P y) ==> adm(%x.!y. P y x)"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   390
by (fast intro: admI elim: admD)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   391
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   392
lemmas adm_all2 = allI [THEN adm_all, standard]
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   393
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   394
lemma adm_subst: "[|cont t; adm P|] ==> adm(%x. P (t x))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   395
apply (rule admI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   396
apply (simplesubst cont2contlub [THEN contlubE, THEN spec, THEN mp])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   397
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   398
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   399
apply (erule admD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   400
apply (erule cont2mono [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   401
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   402
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   403
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   404
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   405
lemma adm_UU_not_less: "adm(%x.~ UU << t(x))"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   406
by simp
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   407
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   408
lemma adm_not_UU: "cont(t)==> adm(%x.~ (t x) = UU)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   409
apply (unfold adm_def)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   410
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   411
apply (rule contrapos_nn)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   412
apply (erule spec)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   413
apply (rule chain_UU_I [THEN spec])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   414
apply (erule cont2mono [THEN ch2ch_monofun])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   415
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   416
apply (erule cont2contlub [THEN contlubE, THEN spec, THEN mp, THEN subst])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   417
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   418
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   419
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   420
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   421
lemma adm_eq: "[|cont u ; cont v|]==> adm(%x. u x = v x)"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   422
by (simp add: po_eq_conv)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   423
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   424
text {* admissibility for disjunction is hard to prove. It takes 10 Lemmas *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   425
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   426
lemma adm_disj_lemma1: "!n. P(Y n)|Q(Y n) ==> (? i.!j. R i j --> Q(Y(j))) | (!i.? j. R i j & P(Y(j)))"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   427
by fast
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   428
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   429
lemma adm_disj_lemma2: "[| adm(Q); ? X. chain(X) & (!n. Q(X(n))) & 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   430
      lub(range(Y))=lub(range(X))|] ==> Q(lub(range(Y)))"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   431
by (force elim: admD)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   432
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   433
lemma adm_disj_lemma3: "chain Y ==> chain (%m. if m < Suc i then Y (Suc i) else Y m)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   434
apply (unfold chain_def)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   435
apply (simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   436
apply safe
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   437
apply (subgoal_tac "ia = i")
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   438
apply (simp_all)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   439
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   440
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   441
lemma adm_disj_lemma4: "!j. i < j --> Q(Y(j))  ==> !n. Q( if n < Suc i then Y(Suc i) else Y n)"
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   442
by (simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   443
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   444
lemma adm_disj_lemma5: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   445
  "!!Y::nat=>'a::cpo. [| chain(Y); ! j. i < j --> Q(Y(j)) |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   446
          lub(range(Y)) = lub(range(%m. if m< Suc(i) then Y(Suc(i)) else Y m))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   447
apply (safe intro!: lub_equal2 adm_disj_lemma3)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   448
prefer 2 apply (assumption)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   449
apply (rule_tac x = "i" in exI)
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   450
apply (simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   451
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   452
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   453
lemma adm_disj_lemma6: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   454
  "[| chain(Y::nat=>'a::cpo); ? i. ! j. i < j --> Q(Y(j)) |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   455
            ? X. chain(X) & (! n. Q(X(n))) & lub(range(Y)) = lub(range(X))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   456
apply (erule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   457
apply (rule_tac x = "%m. if m<Suc (i) then Y (Suc (i)) else Y m" in exI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   458
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   459
apply (rule adm_disj_lemma3)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   460
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   461
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   462
apply (rule adm_disj_lemma4)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   463
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   464
apply (rule adm_disj_lemma5)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   465
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   466
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   467
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   468
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   469
lemma adm_disj_lemma7: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   470
  "[| chain(Y::nat=>'a::cpo); ! i. ? j. i < j & P(Y(j))  |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   471
            chain(%m. Y(Least(%j. m<j & P(Y(j)))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   472
apply (rule chainI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   473
apply (rule chain_mono3)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   474
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   475
apply (rule Least_le)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   476
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   477
apply (rule Suc_lessD)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   478
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   479
apply (erule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   480
apply (rule LeastI [THEN conjunct1])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   481
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   482
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   483
apply (erule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   484
apply (rule LeastI [THEN conjunct2])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   485
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   486
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   487
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   488
lemma adm_disj_lemma8: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   489
  "[| ! i. ? j. i < j & P(Y(j)) |] ==> ! m. P(Y(LEAST j::nat. m<j & P(Y(j))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   490
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   491
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   492
apply (erule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   493
apply (erule LeastI [THEN conjunct2])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   494
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   495
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   496
lemma adm_disj_lemma9: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   497
  "[| chain(Y::nat=>'a::cpo); ! i. ? j. i < j & P(Y(j)) |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   498
            lub(range(Y)) = lub(range(%m. Y(Least(%j. m<j & P(Y(j))))))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   499
apply (rule antisym_less)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   500
apply (rule lub_mono)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   501
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   502
apply (rule adm_disj_lemma7)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   503
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   504
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   505
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   506
apply (rule chain_mono)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   507
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   508
apply (erule allE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   509
apply (erule exE)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   510
apply (rule LeastI [THEN conjunct1])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   511
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   512
apply (rule lub_mono3)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   513
apply (rule adm_disj_lemma7)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   514
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   515
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   516
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   517
apply (intro strip)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   518
apply (rule exI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   519
apply (rule chain_mono)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   520
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   521
apply (rule lessI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   522
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   523
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   524
lemma adm_disj_lemma10: "[| chain(Y::nat=>'a::cpo); ! i. ? j. i < j & P(Y(j)) |] ==> 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   525
            ? X. chain(X) & (! n. P(X(n))) & lub(range(Y)) = lub(range(X))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   526
apply (rule_tac x = "%m. Y (Least (%j. m<j & P (Y (j))))" in exI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   527
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   528
apply (rule adm_disj_lemma7)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   529
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   530
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   531
apply (rule conjI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   532
apply (rule adm_disj_lemma8)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   533
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   534
apply (rule adm_disj_lemma9)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   535
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   536
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   537
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   538
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   539
lemma adm_disj_lemma12: "[| adm(P); chain(Y);? i. ! j. i < j --> P(Y(j))|]==>P(lub(range(Y)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   540
apply (erule adm_disj_lemma2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   541
apply (erule adm_disj_lemma6)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   542
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   543
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   544
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   545
lemma adm_lemma11: 
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   546
"[| adm(P); chain(Y); ! i. ? j. i < j & P(Y(j)) |]==>P(lub(range(Y)))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   547
apply (erule adm_disj_lemma2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   548
apply (erule adm_disj_lemma10)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   549
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   550
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   551
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   552
lemma adm_disj: "[| adm P; adm Q |] ==> adm(%x. P x | Q x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   553
apply (rule admI)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   554
apply (rule adm_disj_lemma1 [THEN disjE])
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   555
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   556
apply (rule disjI2)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   557
apply (erule adm_disj_lemma12)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   558
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   559
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   560
apply (rule disjI1)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   561
apply (erule adm_lemma11)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   562
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   563
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   564
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   565
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   566
lemma adm_imp: "[| adm(%x.~(P x)); adm Q |] ==> adm(%x. P x --> Q x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   567
apply (subgoal_tac " (%x. P x --> Q x) = (%x. ~P x | Q x) ")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   568
apply (erule ssubst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   569
apply (erule adm_disj)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   570
apply assumption
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   571
apply (simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   572
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   573
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   574
lemma adm_iff: "[| adm (%x. P x --> Q x); adm (%x. Q x --> P x) |]  
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   575
            ==> adm (%x. P x = Q x)"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   576
apply (subgoal_tac " (%x. P x = Q x) = (%x. (P x --> Q x) & (Q x --> P x))")
15637
d2a06007ebfa changed comments to text blocks, cleaned up a few proofs
huffman
parents: 15577
diff changeset
   577
apply (simp)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   578
apply (rule ext)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   579
apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   580
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   581
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   582
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   583
lemma adm_not_conj: "[| adm (%x. ~ P x); adm (%x. ~ Q x) |] ==> adm (%x. ~ (P x & Q x))"
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   584
apply (subgoal_tac " (%x. ~ (P x & Q x)) = (%x. ~ P x | ~ Q x) ")
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   585
apply (rule_tac [2] ext)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   586
prefer 2 apply fast
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   587
apply (erule ssubst)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   588
apply (erule adm_disj)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   589
apply assumption
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   590
done
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   591
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   592
lemmas adm_lemmas = adm_not_free adm_imp adm_disj adm_eq adm_not_UU
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   593
        adm_UU_not_less adm_all2 adm_not_less adm_not_conj adm_iff
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   594
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   595
declare adm_lemmas [simp]
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents: 14981
diff changeset
   596
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   597
end
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
   598