src/CCL/Gfp.ML
author wenzelm
Sat, 17 Sep 2005 17:35:26 +0200
changeset 17456 bcf7544875b2
parent 3837 d7f033c74b38
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     1
(*  Title:      CCL/Gfp.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
(*** Proof of Knaster-Tarski Theorem using gfp ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
(* gfp(f) is the least upper bound of {u. u <= f(u)} *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
     9
val prems = goalw (the_context ()) [gfp_def] "[| A <= f(A) |] ==> A <= gfp(f)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
by (rtac (CollectI RS Union_upper) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
by (resolve_tac prems 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    12
qed "gfp_upperbound";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    14
val prems = goalw (the_context ()) [gfp_def]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
    "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
by (REPEAT (ares_tac ([Union_least]@prems) 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
by (etac CollectD 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    18
qed "gfp_least";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    20
val [mono] = goal (the_context ()) "mono(f) ==> gfp(f) <= f(gfp(f))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac,
1459
d12da312eff4 expanded tabs
clasohm
parents: 757
diff changeset
    22
            rtac (mono RS monoD), rtac gfp_upperbound, atac]);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    23
qed "gfp_lemma2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    25
val [mono] = goal (the_context ()) "mono(f) ==> f(gfp(f)) <= gfp(f)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), 
1459
d12da312eff4 expanded tabs
clasohm
parents: 757
diff changeset
    27
            rtac gfp_lemma2, rtac mono]);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    28
qed "gfp_lemma3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    30
val [mono] = goal (the_context ()) "mono(f) ==> gfp(f) = f(gfp(f))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    32
qed "gfp_Tarski";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
(*** Coinduction rules for greatest fixed points ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
(*weak version*)
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    37
val prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
    "[| a: A;  A <= f(A) |] ==> a : gfp(f)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
by (rtac (gfp_upperbound RS subsetD) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
by (REPEAT (ares_tac prems 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    41
qed "coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    43
val [prem,mono] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
    "[| A <= f(A) Un gfp(f);  mono(f) |] ==>  \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
\    A Un gfp(f) <= f(A Un gfp(f))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
by (rtac subset_trans 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
by (rtac (mono RS mono_Un) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
by (rtac (mono RS gfp_Tarski RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
by (rtac (prem RS Un_least) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
by (rtac Un_upper2 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    51
qed "coinduct2_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
(*strong version, thanks to Martin Coen*)
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    54
val ainA::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
    "[| a: A;  A <= f(A) Un gfp(f);  mono(f) |] ==> a : gfp(f)";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    56
by (rtac coinduct 1);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    57
by (rtac (prems MRS coinduct2_lemma) 2);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    58
by (resolve_tac [ainA RS UnI1] 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    59
qed "coinduct2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
(***  Even Stronger version of coinduct  [by Martin Coen]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
         - instead of the condition  A <= f(A)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
                           consider  A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    65
val [prem] = goal (the_context ()) "mono(f) ==> mono(%x. f(x) Un A Un B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    67
qed "coinduct3_mono_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    69
val [prem,mono] = goal (the_context ())
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 2035
diff changeset
    70
    "[| A <= f(lfp(%x. f(x) Un A Un gfp(f)));  mono(f) |] ==> \
d7f033c74b38 fixed dots;
wenzelm
parents: 2035
diff changeset
    71
\    lfp(%x. f(x) Un A Un gfp(f)) <= f(lfp(%x. f(x) Un A Un gfp(f)))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
by (rtac subset_trans 1);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    73
by (rtac (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
by (rtac (Un_least RS Un_least) 1);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    75
by (rtac subset_refl 1);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    76
by (rtac prem 1);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    77
by (rtac (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
by (rtac (mono RS monoD) 1);
2035
e329b36d9136 Ran expandshort; used stac instead of ssubst
paulson
parents: 1459
diff changeset
    79
by (stac (mono RS coinduct3_mono_lemma RS lfp_Tarski) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
by (rtac Un_upper2 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    81
qed "coinduct3_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    83
val ainA::prems = goal (the_context ())
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 2035
diff changeset
    84
    "[| a:A;  A <= f(lfp(%x. f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    85
by (rtac coinduct 1);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    86
by (rtac (prems MRS coinduct3_lemma) 2);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    87
by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1);
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 0
diff changeset
    88
by (rtac (ainA RS UnI2 RS UnI1) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    89
qed "coinduct3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
(** Definition forms of gfp_Tarski, to control unfolding **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    94
val [rew,mono] = goal (the_context ()) "[| h==gfp(f);  mono(f) |] ==> h = f(h)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
by (rtac (mono RS gfp_Tarski) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    97
qed "def_gfp_Tarski";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
    99
val rew::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
    "[| h==gfp(f);  a:A;  A <= f(A) |] ==> a: h";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
by (REPEAT (ares_tac (prems @ [coinduct]) 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   103
qed "def_coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
   105
val rew::prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
    "[| h==gfp(f);  a:A;  A <= f(A) Un h; mono(f) |] ==> a: h";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   109
qed "def_coinduct2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
   111
val rew::prems = goal (the_context ())
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 2035
diff changeset
   112
    "[| h==gfp(f);  a:A;  A <= f(lfp(%x. f(x) Un A Un h)); mono(f) |] ==> a: h";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   115
qed "def_coinduct3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
(*Monotonicity of gfp!*)
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 3837
diff changeset
   118
val prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
    "[| mono(f);  !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
by (rtac gfp_upperbound 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
by (rtac subset_trans 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
by (rtac gfp_lemma2 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
by (resolve_tac prems 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
by (resolve_tac prems 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   125
qed "gfp_mono";