src/ZF/Fixedpt.ML
author paulson
Thu, 17 Jan 2002 12:45:52 +0100
changeset 12789 459b5de466b2
parent 10216 e928bdf62014
permissions -rw-r--r--
new definitions from Sidi Ehmety
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
     1
(*  Title:      ZF/fixedpt.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
     6
Least and greatest fixed points; the Knaster-Tarski Theorem
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Proved in the lattice of subsets of D, namely Pow(D), with Inter as glb
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
(*** Monotone operators ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    13
val prems = Goalw [bnd_mono_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
    "[| h(D)<=D;  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
\       !!W X. [| W<=D;  X<=D;  W<=X |] ==> h(W) <= h(X)  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
\    |] ==> bnd_mono(D,h)";  
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
by (REPEAT (ares_tac (prems@[conjI,allI,impI]) 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
     ORELSE etac subset_trans 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    19
qed "bnd_monoI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    21
Goalw [bnd_mono_def] "bnd_mono(D,h) ==> h(D) <= D";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    22
by (etac conjunct1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    23
qed "bnd_monoD1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    25
Goalw [bnd_mono_def] "[| bnd_mono(D,h);  W<=X;  X<=D |] ==> h(W) <= h(X)";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    26
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    27
qed "bnd_monoD2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
    29
val [major,minor] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
    "[| bnd_mono(D,h);  X<=D |] ==> h(X) <= D";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
by (rtac (major RS bnd_monoD2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
by (rtac (major RS bnd_monoD1) 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
by (rtac minor 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
by (rtac subset_refl 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    35
qed "bnd_mono_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    37
Goal "[| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
\                         h(A) Un h(B) <= h(A Un B)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
     ORELSE etac bnd_monoD2 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    41
qed "bnd_mono_Un";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
(*Useful??*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    44
Goal "[| bnd_mono(D,h);  A <= D;  B <= D |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
\                        h(A Int B) <= h(A) Int h(B)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
     ORELSE etac bnd_monoD2 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    48
qed "bnd_mono_Int";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
(**** Proof of Knaster-Tarski Theorem for the lfp ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
(*lfp is contained in each pre-fixedpoint*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 3016
diff changeset
    53
Goalw [lfp_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    54
    "[| h(A) <= A;  A<=D |] ==> lfp(D,h) <= A";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2469
diff changeset
    55
by (Blast_tac 1);
744
2054fa3c8d76 tidied proofs, using fast_tac etc. as much as possible
lcp
parents: 684
diff changeset
    56
(*or  by (rtac (PowI RS CollectI RS Inter_lower) 1); *)
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    57
qed "lfp_lowerbound";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 3016
diff changeset
    60
Goalw [lfp_def,Inter_def] "lfp(D,h) <= D";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2469
diff changeset
    61
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    62
qed "lfp_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
(*Used in datatype package*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
    65
val [rew] = goal (the_context ()) "A==lfp(D,h) ==> A <= D";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
by (rtac lfp_subset 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    68
qed "def_lfp_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    70
val prems = Goalw [lfp_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
    "[| h(D) <= D;  !!X. [| h(X) <= X;  X<=D |] ==> A<=X |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
\    A <= lfp(D,h)";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    73
by (rtac (Pow_top RS CollectI RS Inter_greatest) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [CollectE,PowD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    75
qed "lfp_greatest";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
    77
val hmono::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
    "[| bnd_mono(D,h);  h(A)<=A;  A<=D |] ==> h(lfp(D,h)) <= A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
by (rtac (hmono RS bnd_monoD2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
by (rtac lfp_lowerbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
by (REPEAT (resolve_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    82
qed "lfp_lemma1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    84
Goal "bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
by (rtac (bnd_monoD1 RS lfp_greatest) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
by (rtac lfp_lemma1 2);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    87
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    88
qed "lfp_lemma2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
    90
val [hmono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
    "bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
by (rtac lfp_lowerbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
by (rtac (hmono RS bnd_monoD2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
by (rtac (hmono RS lfp_lemma2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
by (rtac (hmono RS bnd_mono_subset) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
by (REPEAT (rtac lfp_subset 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
    97
qed "lfp_lemma3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
    99
Goal "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   100
by (REPEAT (ares_tac [equalityI,lfp_lemma2,lfp_lemma3] 1));
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   101
qed "lfp_unfold";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
(*Definition form, to control unfolding*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   104
val [rew,mono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
    "[| A==lfp(D,h);  bnd_mono(D,h) |] ==> A = h(A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
by (rewtac rew);
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   107
by (rtac (mono RS lfp_unfold) 1);
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   108
qed "def_lfp_unfold";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
(*** General induction rule for least fixedpoints ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   112
val [hmono,indstep] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
    "[| bnd_mono(D,h);  !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
\    |] ==> h(Collect(lfp(D,h),P)) <= Collect(lfp(D,h),P)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
by (rtac subsetI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
by (rtac CollectI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
by (etac indstep 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
by (rtac (hmono RS lfp_lemma2 RS subsetD) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
by (rtac (hmono RS bnd_monoD2 RS subsetD) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
by (REPEAT (ares_tac [Collect_subset, lfp_subset] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   121
qed "Collect_is_pre_fixedpt";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
(*This rule yields an induction hypothesis in which the components of a
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
  data structure may be assumed to be elements of lfp(D,h)*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   125
val prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   126
    "[| bnd_mono(D,h);  a : lfp(D,h);                   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   127
\       !!x. x : h(Collect(lfp(D,h),P)) ==> P(x)        \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
\    |] ==> P(a)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   129
by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
by (REPEAT (ares_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   132
qed "induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
(*Definition form, to control unfolding*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   135
val rew::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   136
    "[| A == lfp(D,h);  bnd_mono(D,h);  a:A;   \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
\       !!x. x : h(Collect(A,P)) ==> P(x) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
\    |] ==> P(a)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
by (rtac induct 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   141
qed "def_induct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   142
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
(*This version is useful when "A" is not a subset of D;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   144
  second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   145
val [hsub,hmono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
    "[| h(D Int A) <= A;  bnd_mono(D,h) |] ==> lfp(D,h) <= A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
by (rtac (lfp_lowerbound RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
by (rtac (hmono RS bnd_mono_subset RS Int_greatest) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   149
by (REPEAT (resolve_tac [hsub,Int_lower1,Int_lower2] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   150
qed "lfp_Int_lowerbound";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
(*Monotonicity of lfp, where h precedes i under a domain-like partial order
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   153
  monotonicity of h is not strictly necessary; h must be bounded by D*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   154
val [hmono,imono,subhi] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   155
    "[| bnd_mono(D,h);  bnd_mono(E,i);          \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
\       !!X. X<=D ==> h(X) <= i(X)  |] ==> lfp(D,h) <= lfp(E,i)";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
   157
by (rtac (bnd_monoD1 RS lfp_greatest) 1);
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
   158
by (rtac imono 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
by (rtac (hmono RSN (2, lfp_Int_lowerbound)) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
by (rtac (Int_lower1 RS subhi RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
by (rtac (imono RS bnd_monoD2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
by (REPEAT (ares_tac [Int_lower2] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   163
qed "lfp_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
(*This (unused) version illustrates that monotonicity is not really needed,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
  but both lfp's must be over the SAME set D;  Inter is anti-monotonic!*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   167
val [isubD,subhi] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   168
    "[| i(D) <= D;  !!X. X<=D ==> h(X) <= i(X)  |] ==> lfp(D,h) <= lfp(D,i)";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
   169
by (rtac lfp_greatest 1);
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
   170
by (rtac isubD 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   171
by (rtac lfp_lowerbound 1);
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
   172
by (etac (subhi RS subset_trans) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   174
qed "lfp_mono2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
(**** Proof of Knaster-Tarski Theorem for the gfp ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   179
(*gfp contains each post-fixedpoint that is contained in D*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   180
Goalw [gfp_def] "[| A <= h(A);  A<=D |] ==> A <= gfp(D,h)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
by (rtac (PowI RS CollectI RS Union_upper) 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   182
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   183
qed "gfp_upperbound";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 3016
diff changeset
   185
Goalw [gfp_def] "gfp(D,h) <= D";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2469
diff changeset
   186
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   187
qed "gfp_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
(*Used in datatype package*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   190
val [rew] = goal (the_context ()) "A==gfp(D,h) ==> A <= D";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
by (rtac gfp_subset 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   193
qed "def_gfp_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   195
val hmono::prems = Goalw [gfp_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
    "[| bnd_mono(D,h);  !!X. [| X <= h(X);  X<=D |] ==> X<=A |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
\    gfp(D,h) <= A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
by (fast_tac (subset_cs addIs ((hmono RS bnd_monoD1)::prems)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   199
qed "gfp_least";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   201
val hmono::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
    "[| bnd_mono(D,h);  A<=h(A);  A<=D |] ==> A <= h(gfp(D,h))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
by (rtac (hmono RS bnd_monoD2 RSN (2,subset_trans)) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
by (rtac gfp_subset 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
by (rtac gfp_upperbound 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
by (REPEAT (resolve_tac prems 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   207
qed "gfp_lemma1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   209
Goal "bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
by (rtac gfp_least 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   211
by (rtac gfp_lemma1 2);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   212
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   213
qed "gfp_lemma2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   215
val [hmono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
    "bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
by (rtac gfp_upperbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
by (rtac (hmono RS bnd_monoD2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
by (rtac (hmono RS gfp_lemma2) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   220
by (REPEAT (rtac ([hmono, gfp_subset] MRS bnd_mono_subset) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   221
qed "gfp_lemma3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   222
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   223
Goal "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   224
by (REPEAT (ares_tac [equalityI,gfp_lemma2,gfp_lemma3] 1));
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   225
qed "gfp_unfold";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   226
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
(*Definition form, to control unfolding*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   228
val [rew,mono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
    "[| A==gfp(D,h);  bnd_mono(D,h) |] ==> A = h(A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
by (rewtac rew);
10216
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   231
by (rtac (mono RS gfp_unfold) 1);
e928bdf62014 renamed fp_Tarski to fp_unfold
paulson
parents: 9907
diff changeset
   232
qed "def_gfp_unfold";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   233
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   234
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
(*** Coinduction rules for greatest fixed points ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
(*weak version*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   238
Goal "[| a: X;  X <= h(X);  X <= D |] ==> a : gfp(D,h)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   240
qed "weak_coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   241
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   242
val [subs_h,subs_D,mono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   243
    "[| X <= h(X Un gfp(D,h));  X <= D;  bnd_mono(D,h) |] ==>  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   244
\    X Un gfp(D,h) <= h(X Un gfp(D,h))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
by (rtac (subs_h RS Un_least) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   246
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
by (rtac (Un_upper2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
by (rtac ([mono, subs_D, gfp_subset] MRS bnd_mono_Un) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   249
qed "coinduct_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   250
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   251
(*strong version*)
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   252
Goal "[| bnd_mono(D,h);  a: X;  X <= h(X Un gfp(D,h));  X <= D |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   253
\           a : gfp(D,h)";
647
fb7345cccddc ZF/Fixedpt/coinduct: modified proof to suppress deep unification
lcp
parents: 484
diff changeset
   254
by (rtac weak_coinduct 1);
fb7345cccddc ZF/Fixedpt/coinduct: modified proof to suppress deep unification
lcp
parents: 484
diff changeset
   255
by (etac coinduct_lemma 2);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
by (REPEAT (ares_tac [gfp_subset, UnI1, Un_least] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   257
qed "coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
(*Definition form, to control unfolding*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 5321
diff changeset
   260
val rew::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   261
    "[| A == gfp(D,h);  bnd_mono(D,h);  a: X;  X <= h(X Un A);  X <= D |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
\    a : A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   263
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
by (rtac coinduct 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   266
qed "def_coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   268
(*Lemma used immediately below!*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   269
val [subsA,XimpP] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
    "[| X <= A;  !!z. z:X ==> P(z) |] ==> X <= Collect(A,P)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   271
by (rtac (subsA RS subsetD RS CollectI RS subsetI) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   272
by (assume_tac 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   273
by (etac XimpP 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   274
qed "subset_Collect";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
(*The version used in the induction/coinduction package*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   277
val prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   278
    "[| A == gfp(D, %w. Collect(D,P(w)));  bnd_mono(D, %w. Collect(D,P(w)));  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
\       a: X;  X <= D;  !!z. z: X ==> P(X Un A, z) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   280
\    a : A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
by (rtac def_coinduct 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   282
by (REPEAT (ares_tac (subset_Collect::prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   283
qed "def_Collect_coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   284
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
(*Monotonicity of gfp!*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5147
diff changeset
   286
val [hmono,subde,subhi] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   287
    "[| bnd_mono(D,h);  D <= E;                 \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
\       !!X. X<=D ==> h(X) <= i(X)  |] ==> gfp(D,h) <= gfp(E,i)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   289
by (rtac gfp_upperbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
by (rtac (gfp_subset RS subhi) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
by (rtac ([gfp_subset, subde] MRS subset_trans) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 744
diff changeset
   293
qed "gfp_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294