src/HOL/simpdata.ML
author haftmann
Fri, 13 Oct 2006 12:32:44 +0200
changeset 21009 0eae3fb48936
parent 20973 0b8e436ed071
child 21045 66d6d1b0ddfa
permissions -rw-r--r--
lifted claset setup from ML to Isar level
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1264
diff changeset
     1
(*  Title:      HOL/simpdata.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:     Tobias Nipkow
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5278
diff changeset
     6
Instantiation of the generic simplifier for HOL.
923
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
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
     9
(** tools setup **)
4351
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    10
9851
e22db9397e17 iff declarations moved to clasimp.ML;
wenzelm
parents: 9832
diff changeset
    11
structure Quantifier1 = Quantifier1Fun
e22db9397e17 iff declarations moved to clasimp.ML;
wenzelm
parents: 9832
diff changeset
    12
(struct
4351
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    13
  (*abstract syntax*)
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    14
  fun dest_eq ((c as Const("op =",_)) $ s $ t) = SOME (c, s, t)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15423
diff changeset
    15
    | dest_eq _ = NONE;
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    16
  fun dest_conj ((c as Const("op &",_)) $ s $ t) = SOME (c, s, t)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15423
diff changeset
    17
    | dest_conj _ = NONE;
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    18
  fun dest_imp ((c as Const("op -->",_)) $ s $ t) = SOME (c, s, t)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15423
diff changeset
    19
    | dest_imp _ = NONE;
4351
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    20
  val conj = HOLogic.conj
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    21
  val imp  = HOLogic.imp
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    22
  (*rules*)
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    23
  val iff_reflection = HOL.eq_reflection
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    24
  val iffI = HOL.iffI
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    25
  val iff_trans = HOL.trans
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    26
  val conjI= HOL.conjI
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    27
  val conjE= HOL.conjE
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    28
  val impI = HOL.impI
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    29
  val mp   = HOL.mp
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    30
  val uncurry = thm "uncurry"
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    31
  val exI  = HOL.exI
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
    32
  val exE  = HOL.exE
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    33
  val iff_allI = thm "iff_allI"
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    34
  val iff_exI = thm "iff_exI"
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    35
  val all_comm = thm "all_comm"
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    36
  val ex_comm = thm "ex_comm"
4351
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    37
end);
36b28f78ed1b Tidying and some comments
paulson
parents: 4327
diff changeset
    38
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    39
structure HOL =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    40
struct
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    41
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    42
open HOL;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    43
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    44
val Eq_FalseI = thm "Eq_FalseI";
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    45
val Eq_TrueI = thm "Eq_TrueI";
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    46
val simp_implies_def = thm "simp_implies_def";
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    47
val simp_impliesI = thm "simp_impliesI";
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    48
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    49
fun mk_meta_eq r = r RS eq_reflection;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    50
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    51
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    52
fun mk_eq thm = case concl_of thm
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    53
  (*expects Trueprop if not == *)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    54
  of Const ("==",_) $ _ $ _ => thm
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    55
   | _ $ (Const ("op =", _) $ _ $ _) => mk_meta_eq thm
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    56
   | _ $ (Const ("Not", _) $ _) => thm RS Eq_FalseI
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    57
   | _ => thm RS Eq_TrueI;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    58
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    59
fun mk_eq_True r =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    60
  SOME (r RS meta_eq_to_obj_eq RS Eq_TrueI) handle Thm.THM _ => NONE;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    61
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    62
(* Produce theorems of the form
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    63
  (P1 =simp=> ... =simp=> Pn => x == y) ==> (P1 =simp=> ... =simp=> Pn => x = y)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    64
*)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    65
fun lift_meta_eq_to_obj_eq i st =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    66
  let
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    67
    fun count_imp (Const ("HOL.simp_implies", _) $ P $ Q) = 1 + count_imp Q
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    68
      | count_imp _ = 0;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    69
    val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1)))
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    70
  in if j = 0 then meta_eq_to_obj_eq
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    71
    else
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    72
      let
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    73
        val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    74
        fun mk_simp_implies Q = foldr (fn (R, S) =>
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    75
          Const ("HOL.simp_implies", propT --> propT --> propT) $ R $ S) Q Ps
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    76
        val aT = TFree ("'a", HOLogic.typeS);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    77
        val x = Free ("x", aT);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    78
        val y = Free ("y", aT)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    79
      in Goal.prove_global (Thm.theory_of_thm st) []
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    80
        [mk_simp_implies (Logic.mk_equals (x, y))]
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    81
        (mk_simp_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y))))
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    82
        (fn prems => EVERY
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    83
         [rewrite_goals_tac [simp_implies_def],
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    84
          REPEAT (ares_tac (meta_eq_to_obj_eq :: map (rewrite_rule [simp_implies_def]) prems) 1)])
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    85
      end
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    86
  end;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    87
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    88
(*Congruence rules for = (instead of ==)*)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    89
fun mk_meta_cong rl = zero_var_indexes
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    90
  (let val rl' = Seq.hd (TRYALL (fn i => fn st =>
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    91
     rtac (lift_meta_eq_to_obj_eq i st) i st) rl)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    92
   in mk_meta_eq rl' handle THM _ =>
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    93
     if can Logic.dest_equals (concl_of rl') then rl'
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    94
     else error "Conclusion of congruence rules must be =-equality"
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    95
   end);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    96
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    97
(*
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    98
val mk_atomize:      (string * thm list) list -> thm -> thm list
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
    99
looks too specific to move it somewhere else
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   100
*)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   101
fun mk_atomize pairs =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   102
  let
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   103
    fun atoms thm = case concl_of thm
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   104
     of Const("Trueprop", _) $ p => (case head_of p
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   105
        of Const(a, _) => (case AList.lookup (op =) pairs a
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   106
           of SOME rls => maps atoms ([thm] RL rls)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   107
            | NONE => [thm])
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   108
         | _ => [thm])
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   109
      | _ => [thm]
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   110
  in atoms end;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   111
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   112
fun mksimps pairs =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   113
  (map_filter (try mk_eq) o mk_atomize pairs o gen_all);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   114
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   115
fun unsafe_solver_tac prems =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   116
  (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   117
  FIRST'[resolve_tac(reflexive_thm :: TrueI :: refl :: prems), atac, etac FalseE];
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   118
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   119
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   120
(*No premature instantiation of variables during simplification*)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   121
fun safe_solver_tac prems =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   122
  (fn i => REPEAT_DETERM (match_tac [simp_impliesI] i)) THEN'
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   123
  FIRST'[match_tac(reflexive_thm :: TrueI :: refl :: prems),
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   124
         eq_assume_tac, ematch_tac [FalseE]];
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   125
val safe_solver = mk_solver "HOL safe" safe_solver_tac;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   126
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11220
diff changeset
   127
end;
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11220
diff changeset
   128
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   129
structure SplitterData =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   130
struct
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   131
  structure Simplifier = Simplifier
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   132
  val mk_eq           = HOL.mk_eq
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   133
  val meta_eq_to_iff  = HOL.meta_eq_to_obj_eq
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   134
  val iffD            = HOL.iffD2
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   135
  val disjE           = HOL.disjE
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   136
  val conjE           = HOL.conjE
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   137
  val exE             = HOL.exE
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   138
  val contrapos       = HOL.contrapos_nn
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   139
  val contrapos2      = HOL.contrapos_pp
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   140
  val notnotD         = HOL.notnotD
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   141
end;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   142
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   143
structure Splitter = SplitterFun(SplitterData);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   144
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   145
(* integration of simplifier with classical reasoner *)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   146
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   147
structure Clasimp = ClasimpFun
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   148
 (structure Simplifier = Simplifier and Splitter = Splitter
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   149
  and Classical  = Classical and Blast = Blast
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   150
  val iffD1 = HOL.iffD1 val iffD2 = HOL.iffD2 val notE = HOL.notE);
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12975
diff changeset
   151
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   152
structure HOL =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   153
struct
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   154
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   155
open HOL;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   156
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   157
val mksimps_pairs =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   158
  [("op -->", [mp]), ("op &", [thm "conjunct1", thm "conjunct2"]),
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   159
   ("All", [spec]), ("True", []), ("False", []),
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   160
   ("HOL.If", [thm "if_bool_eq_conj" RS iffD1])];
3913
96e28b16861c New trivial rewrites
paulson
parents: 3904
diff changeset
   161
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   162
val simpset_basic =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   163
  Simplifier.theory_context (the_context ()) empty_ss
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   164
    setsubgoaler asm_simp_tac
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   165
    setSSolver safe_solver
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   166
    setSolver unsafe_solver
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   167
    setmksimps (mksimps mksimps_pairs)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   168
    setmkeqTrue mk_eq_True
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   169
    setmkcong mk_meta_cong;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   170
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   171
fun simplify rews = Simplifier.full_simplify (simpset_basic addsimps rews);
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   172
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   173
fun unfold_tac ths =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   174
  let val ss0 = Simplifier.clear_ss simpset_basic addsimps ths
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   175
  in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   176
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   177
(** simprocs **)
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   178
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   179
(* simproc for proving "(y = x) == False" from premise "~(x = y)" *)
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   180
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   181
val use_neq_simproc = ref true;
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   182
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   183
local
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   184
  val thy = the_context ();
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   185
  val neq_to_EQ_False = thm "not_sym" RS HOL.Eq_FalseI;
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   186
  fun neq_prover sg ss (eq $ lhs $ rhs) =
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   187
    let
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   188
      fun test thm = (case #prop (rep_thm thm) of
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   189
                    _ $ (Not $ (eq' $ l' $ r')) =>
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   190
                      Not = HOLogic.Not andalso eq' = eq andalso
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   191
                      r' aconv lhs andalso l' aconv rhs
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   192
                  | _ => false)
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   193
    in if !use_neq_simproc then case find_first test (prems_of_ss ss)
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   194
     of NONE => NONE
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   195
      | SOME thm => SOME (thm RS neq_to_EQ_False)
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   196
     else NONE
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   197
    end
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   198
in
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   199
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   200
val neq_simproc = Simplifier.simproc thy "neq_simproc" ["x = y"] neq_prover;
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   201
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   202
end; (*local*)
17778
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   203
93d7e524417a changes due to new neq_simproc in simpdata.ML
nipkow
parents: 17325
diff changeset
   204
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   205
(* Simproc for Let *)
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   206
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   207
val use_let_simproc = ref true;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   208
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   209
local
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   210
  val thy = the_context ();
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   211
  val Let_folded = thm "Let_folded";
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   212
  val Let_unfold = thm "Let_unfold";
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   213
  val (f_Let_unfold, x_Let_unfold) =
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   214
      let val [(_$(f$x)$_)] = prems_of Let_unfold
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   215
      in (cterm_of thy f, cterm_of thy x) end
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   216
  val (f_Let_folded, x_Let_folded) =
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   217
      let val [(_$(f$x)$_)] = prems_of Let_folded
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   218
      in (cterm_of thy f, cterm_of thy x) end;
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   219
  val g_Let_folded =
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   220
      let val [(_$_$(g$_))] = prems_of Let_folded in cterm_of thy g end;
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   221
in
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   222
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   223
val let_simproc =
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   224
  Simplifier.simproc thy "let_simp" ["Let x f"]
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   225
   (fn sg => fn ss => fn t =>
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   226
     let val ctxt = Simplifier.the_context ss;
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   227
         val ([t'], ctxt') = Variable.import_terms false [t] ctxt;
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   228
     in Option.map (hd o Variable.export ctxt' ctxt o single)
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   229
      (case t' of (Const ("Let",_)$x$f) => (* x and f are already in normal form *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15423
diff changeset
   230
         if not (!use_let_simproc) then NONE
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   231
         else if is_Free x orelse is_Bound x orelse is_Const x
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   232
         then SOME (thm "Let_def")
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   233
         else
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   234
          let
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   235
             val n = case f of (Abs (x,_,_)) => x | _ => "x";
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   236
             val cx = cterm_of sg x;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   237
             val {T=xT,...} = rep_cterm cx;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   238
             val cf = cterm_of sg f;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   239
             val fx_g = Simplifier.rewrite ss (Thm.capply cf cx);
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   240
             val (_$_$g) = prop_of fx_g;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   241
             val g' = abstract_over (x,g);
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   242
           in (if (g aconv g')
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   243
               then
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   244
                  let
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   245
                    val rl = cterm_instantiate [(f_Let_unfold,cf),(x_Let_unfold,cx)] Let_unfold;
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   246
                  in SOME (rl OF [fx_g]) end
18176
ae9bd644d106 Term.betapply;
wenzelm
parents: 17985
diff changeset
   247
               else if Term.betapply (f,x) aconv g then NONE (*avoid identity conversion*)
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   248
               else let
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   249
                     val abs_g'= Abs (n,xT,g');
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   250
                     val g'x = abs_g'$x;
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   251
                     val g_g'x = symmetric (beta_conversion false (cterm_of sg g'x));
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   252
                     val rl = cterm_instantiate
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   253
                               [(f_Let_folded,cterm_of sg f),(x_Let_folded,cx),
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   254
                                (g_Let_folded,cterm_of sg abs_g')]
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   255
                               Let_folded;
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   256
                   in SOME (rl OF [transitive fx_g g_g'x])
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   257
                   end)
15423
761a4f8e6ad6 added simproc for Let
schirmer
parents: 15184
diff changeset
   258
           end
20014
729a45534001 fixed let-simproc
schirmer
parents: 19472
diff changeset
   259
        | _ => NONE)
20070
3f31fb81b83a let_simproc: activate Variable.import;
wenzelm
parents: 20046
diff changeset
   260
     end)
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   261
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   262
end; (*local*)
1758
60613b065e9b Added ex_imp
nipkow
parents: 1722
diff changeset
   263
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   264
(* A general refutation procedure *)
9713
2c5b42311eb0 cong setup now part of Simplifier;
wenzelm
parents: 9511
diff changeset
   265
5975
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   266
(* Parameters:
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   267
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   268
   test: term -> bool
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   269
   tests if a term is at all relevant to the refutation proof;
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   270
   if not, then it can be discarded. Can improve performance,
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   271
   esp. if disjunctions can be discarded (no case distinction needed!).
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   272
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   273
   prep_tac: int -> tactic
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   274
   A preparation tactic to be applied to the goal once all relevant premises
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   275
   have been moved to the conclusion.
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   276
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   277
   ref_tac: int -> tactic
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   278
   the actual refutation tactic. Should be able to deal with goals
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   279
   [| A1; ...; An |] ==> False
9876
wenzelm
parents: 9875
diff changeset
   280
   where the Ai are atomic, i.e. no top-level &, | or EX
5975
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   281
*)
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   282
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 14749
diff changeset
   283
local
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 14749
diff changeset
   284
  val nnf_simpset =
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   285
    empty_ss setmkeqTrue mk_eq_True
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   286
    setmksimps (mksimps mksimps_pairs)
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   287
    addsimps [thm "imp_conv_disj", thm "iff_conv_conj_imp", thm "de_Morgan_disj", thm "de_Morgan_conj",
20944
34b2c1bb7178 cleanup basic HOL bootstrap
haftmann
parents: 20932
diff changeset
   288
      thm "not_all", thm "not_ex", thm "not_not"];
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   289
  fun prem_nnf_tac i st =
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   290
    full_simp_tac (Simplifier.theory_context (Thm.theory_of_thm st) nnf_simpset) i st;
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 14749
diff changeset
   291
in
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 14749
diff changeset
   292
fun refute_tac test prep_tac ref_tac =
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 14749
diff changeset
   293
  let val refute_prems_tac =
12475
18ba10cc782f Removed pointless backtracking from arith_tac
nipkow
parents: 12281
diff changeset
   294
        REPEAT_DETERM
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   295
              (eresolve_tac [conjE, exE] 1 ORELSE
5975
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   296
               filter_prems_tac test 1 ORELSE
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   297
               etac disjE 1) THEN
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   298
        ((etac notE 1 THEN eq_assume_tac 1) ORELSE
11200
f43fa07536c0 arith_tac now copes with propositional reasoning as well.
nipkow
parents: 11162
diff changeset
   299
         ref_tac 1);
5975
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   300
  in EVERY'[TRY o filter_prems_tac test,
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   301
            REPEAT_DETERM o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac,
5975
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   302
            SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]
cd19eaa90f45 Added a general refutation tactic which works by putting things into nnf first.
nipkow
parents: 5552
diff changeset
   303
  end;
20973
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   304
end; (*local*)
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   305
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   306
val defALL_regroup =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   307
  Simplifier.simproc (the_context ())
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   308
    "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   309
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   310
val defEX_regroup =
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   311
  Simplifier.simproc (the_context ())
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   312
    "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex;
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   313
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   314
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   315
val simpset_simprocs = simpset_basic
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   316
  addsimprocs [defALL_regroup, defEX_regroup, neq_simproc, let_simproc]
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   317
0b8e436ed071 cleaned up HOL bootstrap
haftmann
parents: 20944
diff changeset
   318
end; (*struct*)