src/Provers/quantifier1.ML
author paulson
Sat, 10 Jan 1998 17:59:32 +0100
changeset 4552 bb8ff763c93d
parent 4319 afb60b8bf15e
child 7951 b36913c35699
permissions -rw-r--r--
Simplified proofs by omitting PA = {|XA, ...|} from RA2
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     1
(*  Title:      Provers/quantifier1
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     4
    Copyright   1997  TU Munich
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     5
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     6
Simplification procedures for turning
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     7
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     8
            ? x. ... & x = t & ...
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     9
     into   ? x. x = t & ... & ...
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    10
     where the `? x. x = t &' in the latter formula is eliminated
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    11
           by ordinary simplification. 
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    12
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    13
     and   ! x. (... & x = t & ...) --> P x
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    14
     into  ! x. x = t --> (... & ...) --> P x
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    15
     where the `!x. x=t -->' in the latter formula is eliminated
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    16
           by ordinary simplification.
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    17
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    18
     NB Simproc is only triggered by "!x. P(x) & P'(x) --> Q(x)";
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    19
        "!x. x=t --> P(x)" is covered by the congreunce rule for -->;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    20
        "!x. t=x --> P(x)" must be taken care of by an ordinary rewrite rule.
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    21
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    22
Gries etc call this the "1 point rules"
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    23
*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    24
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    25
signature QUANTIFIER1_DATA =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    26
sig
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    27
  (*abstract syntax*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    28
  val dest_eq: term -> (term*term*term)option
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    29
  val dest_conj: term -> (term*term*term)option
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    30
  val conj: term
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    31
  val imp:  term
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    32
  (*rules*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    33
  val iff_reflection: thm (* P <-> Q ==> P == Q *)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    34
  val iffI:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    35
  val sym:   thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    36
  val conjI: thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    37
  val conjE: thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    38
  val impI:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    39
  val impE:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    40
  val mp:    thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    41
  val exI:   thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    42
  val exE:   thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    43
  val allI:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    44
  val allE:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    45
end;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    46
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    47
signature QUANTIFIER1 =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    48
sig
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    49
  val rearrange_all: Sign.sg -> thm list -> term -> thm option
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    50
  val rearrange_ex:  Sign.sg -> thm list -> term -> thm option
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    51
end;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    52
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    53
functor Quantifier1Fun(Data: QUANTIFIER1_DATA): QUANTIFIER1 =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    54
struct
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    55
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    56
open Data;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    57
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    58
fun def eq = case dest_eq eq of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    59
      Some(c,s,t) =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    60
        if s = Bound 0 andalso not(loose_bvar1(t,0)) then Some eq else
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    61
        if t = Bound 0 andalso not(loose_bvar1(s,0)) then Some(c$t$s)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    62
        else None
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    63
    | None => None;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    64
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    65
fun extract conj = case dest_conj conj of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    66
      Some(conj,P,Q) =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    67
        (case def P of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    68
           Some eq => Some(eq,Q)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    69
         | None =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    70
             (case def Q of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    71
                Some eq => Some(eq,P)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    72
              | None =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    73
                 (case extract P of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    74
                    Some(eq,P') => Some(eq, conj $ P' $ Q)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    75
                  | None =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    76
                      (case extract Q of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    77
                         Some(eq,Q') => Some(eq,conj $ P $ Q')
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    78
                       | None => None))))
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    79
    | None => None;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    80
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    81
fun prove_conv tac sg tu =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    82
  let val meta_eq = cterm_of sg (Logic.mk_equals tu)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    83
  in prove_goalw_cterm [] meta_eq (K [rtac iff_reflection 1, tac])
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    84
     handle ERROR =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    85
            error("The error(s) above occurred while trying to prove " ^
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    86
                  string_of_cterm meta_eq)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    87
  end;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    88
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    89
val prove_all_tac = EVERY1[rtac iffI,
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    90
                       rtac allI, etac allE, rtac impI, rtac impI, etac mp,
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    91
                          REPEAT o (etac conjE),
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    92
                          REPEAT o (ares_tac [conjI] ORELSE' etac sym),
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    93
                       rtac allI, etac allE, rtac impI, REPEAT o (etac conjE),
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    94
                          etac impE, atac ORELSE' etac sym, etac mp,
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    95
                          REPEAT o (ares_tac [conjI])];
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    96
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    97
fun rearrange_all sg _ (F as all $ Abs(x,T,(* --> *)_ $ P $ Q)) =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    98
     (case extract P of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    99
        None => None
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   100
      | Some(eq,P') =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   101
          let val R = imp $ eq $ (imp $ P' $ Q)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   102
          in Some(prove_conv prove_all_tac sg (F,all$Abs(x,T,R))) end)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   103
  | rearrange_all _ _ _ = None;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   104
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   105
val prove_ex_tac = rtac iffI 1 THEN
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   106
    ALLGOALS(EVERY'[etac exE, REPEAT o (etac conjE),
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   107
                    rtac exI, REPEAT o (ares_tac [conjI] ORELSE' etac sym)]);
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   108
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   109
fun rearrange_ex sg _ (F as ex $ Abs(x,T,P)) =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   110
     (case extract P of
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   111
        None => None
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   112
      | Some(eq,Q) =>
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   113
          Some(prove_conv prove_ex_tac sg (F,ex $ Abs(x,T,conj$eq$Q))))
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   114
  | rearrange_ex _ _ _ = None;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   115
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   116
end;