src/HOL/Lfp.ML
author paulson
Fri, 31 Jan 1997 17:13:19 +0100
changeset 2572 8a47f85e7a03
parent 1873 b07ee188f061
child 3842 b55686a7b22c
permissions -rw-r--r--
ex_impE was incorrectly listed as Safe
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/lfp.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
For lfp.thy.  The Knaster-Tarski Theorem
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
open Lfp;
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
(*** Proof of Knaster-Tarski Theorem ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
by (rtac (CollectI RS Inter_lower) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
by (resolve_tac prems 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
qed "lfp_lowerbound";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
val prems = goalw Lfp.thy [lfp_def]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
    "[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
by (REPEAT (ares_tac ([Inter_greatest]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
by (etac CollectD 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
qed "lfp_greatest";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
by (EVERY1 [rtac lfp_greatest, rtac subset_trans,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    28
            rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
qed "lfp_lemma2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD), 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    33
            rtac lfp_lemma2, rtac mono]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
qed "lfp_lemma3";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
qed "lfp_Tarski";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
(*** General induction rule for least fixed points ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
val [lfp,mono,indhyp] = goal Lfp.thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    43
    "[| a: lfp(f);  mono(f);                            \
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    44
\       !!x. [| x: f(lfp(f) Int {x.P(x)}) |] ==> P(x)   \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
\    |] ==> P(a)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
by (res_inst_tac [("a","a")] (Int_lower2 RS subsetD RS CollectD) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
by (rtac (lfp RSN (2, lfp_lowerbound RS subsetD)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
by (EVERY1 [rtac Int_greatest, rtac subset_trans, 
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    49
            rtac (Int_lower1 RS (mono RS monoD)),
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    50
            rtac (mono RS lfp_lemma2),
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    51
            rtac (CollectI RS subsetI), rtac indhyp, atac]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
qed "induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
1746
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1552
diff changeset
    54
bind_thm
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1552
diff changeset
    55
  ("induct2",
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1552
diff changeset
    56
   Prod_Syntax.split_rule
f0c6aabc6c02 Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents: 1552
diff changeset
    57
     (read_instantiate [("a","(a,b)")] induct));
1114
c8dfb56a7e95 Prod is now a parent of Lfp.
nipkow
parents: 923
diff changeset
    58
1125
13a3df2adbe5 Added Park induction to Lfp.
nipkow
parents: 1114
diff changeset
    59
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
(** Definition forms of lfp_Tarski and induct, to control unfolding **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
val [rew,mono] = goal Lfp.thy "[| h==lfp(f);  mono(f) |] ==> h = f(h)";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
by (rewtac rew);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
by (rtac (mono RS lfp_Tarski) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
qed "def_lfp_Tarski";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
val rew::prems = goal Lfp.thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    68
    "[| A == lfp(f);  mono(f);   a:A;                   \
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    69
\       !!x. [| x: f(A Int {x.P(x)}) |] ==> P(x)        \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
\    |] ==> P(a)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    71
by (EVERY1 [rtac induct,        (*backtracking to force correct induction*)
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    72
            REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
qed "def_induct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
(*Monotonicity of lfp!*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
val [prem] = goal Lfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    77
by (rtac (lfp_lowerbound RS lfp_greatest) 1);
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
    78
by (etac (prem RS subset_trans) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
qed "lfp_mono";