src/HOL/Import/shuffler.ML
author haftmann
Fri, 24 Feb 2012 22:46:16 +0100
changeset 46663 7fe029e818c2
parent 46218 ecf6375e2abb
child 46803 f8875c15cbe1
permissions -rw-r--r--
explicit is better than implicit
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14518
diff changeset
     1
(*  Title:      HOL/Import/shuffler.ML
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     2
    Author:     Sebastian Skalberg, TU Muenchen
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     3
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     4
Package for proving two terms equal by normalizing (hence the
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     5
"shuffler" name).  Uses the simplifier for the normalization.
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     6
*)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     7
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     8
signature Shuffler =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     9
sig
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    10
    val debug      : bool Unsynchronized.ref
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    11
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    12
    val norm_term  : theory -> term -> thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    13
    val make_equal : theory -> term -> term -> thm option
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    14
    val set_prop   : theory -> term -> (string * thm) list -> (string * thm) option
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    15
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    16
    val find_potential: theory -> term -> (string * thm) list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    17
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
    18
    val gen_shuffle_tac: Proof.context -> bool -> (string * thm) list -> int -> tactic
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
    19
    val shuffle_tac: Proof.context -> thm list -> int -> tactic
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
    20
    val search_tac : Proof.context -> int -> tactic
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    21
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    22
    val print_shuffles: theory -> unit
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    23
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    24
    val add_shuffle_rule: thm -> theory -> theory
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    25
    val shuffle_attr: attribute
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    26
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
    27
    val setup      : theory -> theory
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    28
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    29
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    30
structure Shuffler :> Shuffler =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    31
struct
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    32
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    33
val debug = Unsynchronized.ref false
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    34
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    35
fun if_debug f x = if !debug then f x else ()
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    36
val message = if_debug writeln
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    37
37146
f652333bbf8e renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
wenzelm
parents: 36945
diff changeset
    38
val string_of_thm = Print_Mode.setmp [] Display.string_of_thm_without_context;
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    39
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33317
diff changeset
    40
structure ShuffleData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    41
(
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    42
  type T = thm list
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    43
  val empty = []
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    44
  val extend = I
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33317
diff changeset
    45
  val merge = Thm.merge_thms
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    46
)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    47
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    48
fun print_shuffles thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    49
  Pretty.writeln (Pretty.big_list "Shuffle theorems:"
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31945
diff changeset
    50
    (map (Display.pretty_thm_global thy) (ShuffleData.get thy)))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    51
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    52
val weaken =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    53
    let
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 26277
diff changeset
    54
        val cert = cterm_of Pure.thy
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    55
        val P = Free("P",propT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    56
        val Q = Free("Q",propT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    57
        val PQ = Logic.mk_implies(P,Q)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    58
        val PPQ = Logic.mk_implies(P,PQ)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    59
        val cP = cert P
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    60
        val cQ = cert Q
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    61
        val cPQ = cert PQ
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    62
        val cPPQ = cert PPQ
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    63
        val th1 = Thm.assume cPQ |> implies_intr_list [cPQ,cP]
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    64
        val th3 = Thm.assume cP
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    65
        val th4 = implies_elim_list (Thm.assume cPPQ) [th3,th3]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    66
                                    |> implies_intr_list [cPPQ,cP]
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    67
    in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    68
        Thm.equal_intr th4 th1 |> Drule.export_without_context
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    69
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    70
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    71
val imp_comm =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    72
    let
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 26277
diff changeset
    73
        val cert = cterm_of Pure.thy
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    74
        val P = Free("P",propT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    75
        val Q = Free("Q",propT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    76
        val R = Free("R",propT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    77
        val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    78
        val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    79
        val cP = cert P
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    80
        val cQ = cert Q
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    81
        val cPQR = cert PQR
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    82
        val cQPR = cert QPR
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    83
        val th1 = implies_elim_list (Thm.assume cPQR) [Thm.assume cP,Thm.assume cQ]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    84
                                    |> implies_intr_list [cPQR,cQ,cP]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    85
        val th2 = implies_elim_list (Thm.assume cQPR) [Thm.assume cQ,Thm.assume cP]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    86
                                    |> implies_intr_list [cQPR,cP,cQ]
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    87
    in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
    88
        Thm.equal_intr th1 th2 |> Drule.export_without_context
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    89
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    90
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    91
val all_comm =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    92
    let
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 26277
diff changeset
    93
        val cert = cterm_of Pure.thy
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    94
        val xT = TFree("'a",[])
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    95
        val yT = TFree("'b",[])
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
    96
        val x = Free("x",xT)
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
    97
        val y = Free("y",yT)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
    98
        val P = Free("P",xT-->yT-->propT)
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
    99
        val lhs = Logic.all x (Logic.all y (P $ x $ y))
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   100
        val rhs = Logic.all y (Logic.all x (P $ x $ y))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   101
        val cl = cert lhs
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   102
        val cr = cert rhs
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   103
        val cx = cert x
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   104
        val cy = cert y
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   105
        val th1 = Thm.assume cr
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   106
                         |> forall_elim_list [cy,cx]
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   107
                         |> forall_intr_list [cx,cy]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   108
                         |> Thm.implies_intr cr
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   109
        val th2 = Thm.assume cl
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   110
                         |> forall_elim_list [cx,cy]
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   111
                         |> forall_intr_list [cy,cx]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   112
                         |> Thm.implies_intr cl
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   113
    in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   114
        Thm.equal_intr th1 th2 |> Drule.export_without_context
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   115
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   116
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   117
val equiv_comm =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   118
    let
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 26277
diff changeset
   119
        val cert = cterm_of Pure.thy
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   120
        val T    = TFree("'a",[])
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   121
        val t    = Free("t",T)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   122
        val u    = Free("u",T)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   123
        val ctu  = cert (Logic.mk_equals(t,u))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   124
        val cut  = cert (Logic.mk_equals(u,t))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   125
        val th1  = Thm.assume ctu |> Thm.symmetric |> Thm.implies_intr ctu
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   126
        val th2  = Thm.assume cut |> Thm.symmetric |> Thm.implies_intr cut
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   127
    in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   128
        Thm.equal_intr th1 th2 |> Drule.export_without_context
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   129
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   130
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   131
(* This simplification procedure rewrites !!x y. P x y
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   132
deterministicly, in order for the normalization function, defined
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   133
below, to handle nested quantifiers robustly *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   134
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   135
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   136
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   137
exception RESULT of int
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   138
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   139
fun find_bound n (Bound i) = if i = n then raise RESULT 0
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   140
                             else if i = n+1 then raise RESULT 1
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   141
                             else ()
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   142
  | find_bound n (t $ u) = (find_bound n t; find_bound n u)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   143
  | find_bound n (Abs(_,_,t)) = find_bound (n+1) t
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   144
  | find_bound _ _ = ()
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   145
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   146
fun swap_bound n (Bound i) = if i = n then Bound (n+1)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   147
                             else if i = n+1 then Bound n
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   148
                             else Bound i
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   149
  | swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   150
  | swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   151
  | swap_bound n t = t
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   152
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   153
fun rew_th thy (xv as (x,xT)) (yv as (y,yT)) t =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   154
    let
46218
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 46201
diff changeset
   155
        val lhs = Logic.list_all ([xv,yv],t)
ecf6375e2abb renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
wenzelm
parents: 46201
diff changeset
   156
        val rhs = Logic.list_all ([yv,xv],swap_bound 0 t)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   157
        val rew = Logic.mk_equals (lhs,rhs)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   158
        val init = Thm.trivial (cterm_of thy rew)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   159
    in
37778
87b5dfe00387 do not intercept ML exceptions -- printing exception positions/text is the job of the Isar/ML toplevel;
wenzelm
parents: 37146
diff changeset
   160
        all_comm RS init
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   161
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   162
46201
afdc69f5156e eliminated dead code;
wenzelm
parents: 45624
diff changeset
   163
fun quant_rewrite thy _ (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   164
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   165
        val res = (find_bound 0 body;2) handle RESULT i => i
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   166
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   167
        case res of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   168
            0 => SOME (rew_th thy (x,xT) (y,yT) body)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   169
          | 1 => if string_ord(y,x) = LESS
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   170
                 then
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   171
                     let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   172
                         val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   173
                         val t_th    = Thm.reflexive (cterm_of thy t)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   174
                         val newt_th = Thm.reflexive (cterm_of thy newt)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   175
                     in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   176
                         SOME (Thm.transitive t_th newt_th)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   177
                     end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   178
                 else NONE
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   179
          | _ => error "norm_term (quant_rewrite) internal error"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   180
     end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14854
diff changeset
   181
  | quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   182
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   183
fun freeze_thaw_term t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   184
    let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43333
diff changeset
   185
        val tvars = Misc_Legacy.term_tvars t
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43333
diff changeset
   186
        val tfree_names = Misc_Legacy.add_term_tfree_names(t,[])
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   187
        val (type_inst,_) =
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33043
diff changeset
   188
            fold (fn (w as (v,_), S) => fn (inst, used) =>
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   189
                      let
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42368
diff changeset
   190
                          val v' = singleton (Name.variant_list used) v
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   191
                      in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   192
                          ((w,TFree(v',S))::inst,v'::used)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   193
                      end)
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33043
diff changeset
   194
                  tvars ([], tfree_names)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   195
        val t' = subst_TVars type_inst t
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   196
    in
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33043
diff changeset
   197
        (t', map (fn (w,TFree(v,S)) => (v,TVar(w,S))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   198
                  | _ => error "Internal error in Shuffler.freeze_thaw") type_inst)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   199
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   200
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   201
fun inst_tfrees thy [] thm = thm
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   202
  | inst_tfrees thy ((name,U)::rest) thm =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   203
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   204
        val cU = ctyp_of thy U
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43333
diff changeset
   205
        val tfrees = Misc_Legacy.add_term_tfrees (prop_of thm,[])
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35408
diff changeset
   206
        val (rens, thm') = Thm.varifyT_global'
20951
868120282837 gen_rem(s) abandoned in favour of remove / subtract
haftmann
parents: 20897
diff changeset
   207
    (remove (op = o apsnd fst) name tfrees) thm
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   208
        val mid =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   209
            case rens of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   210
                [] => thm'
43333
2bdec7f430d3 renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents: 43324
diff changeset
   211
              | [((_, S), idx)] => Drule.instantiate_normalize
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   212
            ([(ctyp_of thy (TVar (idx, S)), cU)], []) thm'
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   213
              | _ => error "Shuffler.inst_tfrees internal error"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   214
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   215
        inst_tfrees thy rest mid
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   216
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   217
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   218
fun is_Abs (Abs _) = true
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   219
  | is_Abs _ = false
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   220
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   221
fun eta_redex (t $ Bound 0) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   222
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   223
        fun free n (Bound i) = i = n
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   224
          | free n (t $ u) = free n t orelse free n u
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   225
          | free n (Abs(_,_,t)) = free (n+1) t
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   226
          | free n _ = false
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   227
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   228
        not (free 0 t)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   229
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   230
  | eta_redex _ = false
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   231
46201
afdc69f5156e eliminated dead code;
wenzelm
parents: 45624
diff changeset
   232
fun eta_contract thy _ origt =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   233
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   234
        val (typet,Tinst) = freeze_thaw_term origt
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   235
        val (init,thaw) = Drule.legacy_freeze_thaw (Thm.reflexive (cterm_of thy typet))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   236
        val final = inst_tfrees thy Tinst o thaw
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   237
        val t = #1 (Logic.dest_equals (prop_of init))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   238
        val _ =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   239
            let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   240
                val lhs = #1 (Logic.dest_equals (prop_of (final init)))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   241
            in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   242
                if not (lhs aconv origt)
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   243
                then
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   244
                  writeln (cat_lines
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   245
                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   246
                      Syntax.string_of_term_global thy origt,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   247
                      Syntax.string_of_term_global thy lhs,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   248
                      Syntax.string_of_term_global thy typet,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   249
                      Syntax.string_of_term_global thy t] @
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   250
                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   251
                else ()
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   252
            end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   253
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   254
        case t of
46201
afdc69f5156e eliminated dead code;
wenzelm
parents: 45624
diff changeset
   255
            Const("all",_) $ (Abs(x,xT,Const("==",_) $ P $ Q)) =>
37778
87b5dfe00387 do not intercept ML exceptions -- printing exception positions/text is the job of the Isar/ML toplevel;
wenzelm
parents: 37146
diff changeset
   256
            (if eta_redex P andalso eta_redex Q
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   257
              then
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   258
                  let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   259
                      val cert = cterm_of thy
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42368
diff changeset
   260
                      val v = Free (singleton (Name.variant_list (Term.add_free_names t [])) "v", xT)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   261
                      val cv = cert v
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   262
                      val ct = cert t
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   263
                      val th = (Thm.assume ct)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   264
                                   |> Thm.forall_elim cv
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   265
                                   |> Thm.abstract_rule x cv
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   266
                      val ext_th = Thm.eta_conversion (cert (Abs(x,xT,P)))
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   267
                      val th' = Thm.transitive (Thm.symmetric ext_th) th
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   268
                      val cu = cert (prop_of th')
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   269
                      val uth = Thm.combination (Thm.assume cu) (Thm.reflexive cv)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   270
                      val uth' = (Thm.beta_conversion false (cert (Abs(x,xT,Q) $ v)))
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   271
                                     |> Thm.transitive uth
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   272
                                     |> Thm.forall_intr cv
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   273
                                     |> Thm.implies_intr cu
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   274
                      val rew_th = Thm.equal_intr (th' |> Thm.implies_intr ct) uth'
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   275
                      val res = final rew_th
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   276
                  in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   277
                       SOME res
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   278
                  end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   279
              else NONE)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   280
          | _ => NONE
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17188
diff changeset
   281
       end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   282
46201
afdc69f5156e eliminated dead code;
wenzelm
parents: 45624
diff changeset
   283
fun eta_expand thy _ origt =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   284
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   285
        val (typet,Tinst) = freeze_thaw_term origt
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   286
        val (init,thaw) = Drule.legacy_freeze_thaw (Thm.reflexive (cterm_of thy typet))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   287
        val final = inst_tfrees thy Tinst o thaw
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   288
        val t = #1 (Logic.dest_equals (prop_of init))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   289
        val _ =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   290
            let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   291
                val lhs = #1 (Logic.dest_equals (prop_of (final init)))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   292
            in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   293
                if not (lhs aconv origt)
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   294
                then
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   295
                  writeln (cat_lines
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   296
                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   297
                      Syntax.string_of_term_global thy origt,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   298
                      Syntax.string_of_term_global thy lhs,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   299
                      Syntax.string_of_term_global thy typet,
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   300
                      Syntax.string_of_term_global thy t] @
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   301
                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   302
                else ()
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   303
            end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   304
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   305
        case t of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   306
            Const("==",T) $ P $ Q =>
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   307
            if is_Abs P orelse is_Abs Q
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   308
            then (case domain_type T of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   309
                      Type("fun",[aT,bT]) =>
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   310
                      let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   311
                          val cert = cterm_of thy
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42368
diff changeset
   312
                          val vname = singleton (Name.variant_list (Term.add_free_names t [])) "v"
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   313
                          val v = Free(vname,aT)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   314
                          val cv = cert v
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   315
                          val ct = cert t
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   316
                          val th1 = (Thm.combination (Thm.assume ct) (Thm.reflexive cv))
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   317
                                        |> Thm.forall_intr cv
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   318
                                        |> Thm.implies_intr ct
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   319
                          val concl = cert (concl_of th1)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   320
                          val th2 = (Thm.assume concl)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   321
                                        |> Thm.forall_elim cv
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   322
                                        |> Thm.abstract_rule vname cv
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   323
                          val (lhs,rhs) = Logic.dest_equals (prop_of th2)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   324
                          val elhs = Thm.eta_conversion (cert lhs)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   325
                          val erhs = Thm.eta_conversion (cert rhs)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   326
                          val th2' = Thm.transitive
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   327
                                         (Thm.transitive (Thm.symmetric elhs) th2)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   328
                                         erhs
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   329
                          val res = Thm.equal_intr th1 (th2' |> Thm.implies_intr concl)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   330
                          val res' = final res
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   331
                      in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   332
                          SOME res'
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   333
                      end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   334
                    | _ => NONE)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   335
            else NONE
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   336
          | _ => error ("Bad eta_expand argument" ^ Syntax.string_of_term_global thy t)
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   337
    end;
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   338
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14848
diff changeset
   339
fun mk_tfree s = TFree("'"^s,[])
20326
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   340
fun mk_free s t = Free (s,t)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   341
val xT = mk_tfree "a"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   342
val yT = mk_tfree "b"
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   343
val x = Free ("x", xT)
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   344
val y = Free ("y", yT)
20326
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   345
val P  = mk_free "P" (xT-->yT-->propT)
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   346
val Q  = mk_free "Q" (xT-->yT)
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   347
val R  = mk_free "R" (xT-->yT)
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   348
val S  = mk_free "S" xT
cbf31171c147 fixed generator
obua
parents: 20224
diff changeset
   349
val S'  = mk_free "S'" xT
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   350
in
17188
a26a4fc323ed Updated import.
obua
parents: 16428
diff changeset
   351
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
   352
fun quant_simproc thy = Simplifier.simproc_global_i
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   353
                           thy
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   354
                           "Ordered rewriting of nested quantifiers"
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   355
                           [Logic.all x (Logic.all y (P $ x $ y))]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   356
                           quant_rewrite
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
   357
fun eta_expand_simproc thy = Simplifier.simproc_global_i
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   358
                         thy
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   359
                         "Smart eta-expansion by equivalences"
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   360
                         [Logic.mk_equals(Q,R)]
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   361
                         eta_expand
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38558
diff changeset
   362
fun eta_contract_simproc thy = Simplifier.simproc_global_i
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   363
                         thy
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   364
                         "Smart handling of eta-contractions"
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   365
                         [Logic.all x (Logic.mk_equals (Q $ x, R $ x))]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   366
                         eta_contract
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   367
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   368
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   369
(* Disambiguates the names of bound variables in a term, returning t
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   370
== t' where all the names of bound variables in t' are unique *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   371
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   372
fun disamb_bound thy t =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   373
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   374
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   375
        fun F (t $ u,idx) =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   376
            let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   377
                val (t',idx') = F (t,idx)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   378
                val (u',idx'') = F (u,idx')
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   379
            in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   380
                (t' $ u',idx'')
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   381
            end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   382
          | F (Abs(x,xT,t),idx) =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   383
            let
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41490
diff changeset
   384
                val x' = "x" ^ string_of_int idx
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   385
                val (t',idx') = F (t,idx+1)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   386
            in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   387
                (Abs(x',xT,t'),idx')
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   388
            end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   389
          | F arg = arg
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   390
        val (t',_) = F (t,0)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   391
        val ct = cterm_of thy t
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   392
        val ct' = cterm_of thy t'
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   393
        val res = Thm.transitive (Thm.reflexive ct) (Thm.reflexive ct')
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   394
        val _ = message ("disamb_term: " ^ (string_of_thm res))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   395
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   396
        res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   397
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   398
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   399
(* Transforms a term t to some normal form t', returning the theorem t
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   400
== t'.  This is originally a help function for make_equal, but might
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   401
be handy in its own right, for example for indexing terms. *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   402
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   403
fun norm_term thy t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   404
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   405
        val norms = ShuffleData.get thy
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 35021
diff changeset
   406
        val ss = Simplifier.global_context thy empty_ss
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   407
          addsimps (map (Thm.transfer thy) norms)
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   408
          addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy]
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   409
        fun chain f th =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   410
            let
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22846
diff changeset
   411
                val rhs = Thm.rhs_of th
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   412
            in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   413
                Thm.transitive th (f rhs)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   414
            end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   415
        val th =
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   416
            t |> disamb_bound thy
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   417
              |> chain (Simplifier.full_rewrite ss)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   418
              |> chain Thm.eta_conversion
36614
b6c031ad3690 minor tuning of Thm.strip_shyps -- no longer pervasive;
wenzelm
parents: 36543
diff changeset
   419
              |> Thm.strip_shyps
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   420
        val _ = message ("norm_term: " ^ (string_of_thm th))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   421
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   422
        th
17463
e9c1574d0caf removed spurious PolyML.exception_trace;
wenzelm
parents: 17440
diff changeset
   423
    end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   424
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   425
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   426
(* Closes a theorem with respect to free and schematic variables (does
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   427
not touch type variables, though). *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   428
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   429
fun close_thm th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   430
    let
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21588
diff changeset
   431
        val thy = Thm.theory_of_thm th
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   432
        val c = prop_of th
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43333
diff changeset
   433
        val vars = Misc_Legacy.add_term_frees (c, Misc_Legacy.add_term_vars(c,[]))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   434
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   435
        Drule.forall_intr_list (map (cterm_of thy) vars) th
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   436
    end
37778
87b5dfe00387 do not intercept ML exceptions -- printing exception positions/text is the job of the Isar/ML toplevel;
wenzelm
parents: 37146
diff changeset
   437
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   438
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   439
(* Normalizes a theorem's conclusion using norm_term. *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   440
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   441
fun norm_thm thy th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   442
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   443
        val c = prop_of th
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   444
    in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   445
        Thm.equal_elim (norm_term thy c) th
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   446
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   447
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   448
(* make_equal thy t u tries to construct the theorem t == u under the
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   449
signature thy.  If it succeeds, SOME (t == u) is returned, otherwise
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14854
diff changeset
   450
NONE is returned. *)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   451
21078
101aefd61aac slight cleanup
haftmann
parents: 20951
diff changeset
   452
fun make_equal thy t u =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   453
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   454
        val t_is_t' = norm_term thy t
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   455
        val u_is_u' = norm_term thy u
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   456
        val th = Thm.transitive t_is_t' (Thm.symmetric u_is_u')
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   457
        val _ = message ("make_equal: SOME " ^ (string_of_thm th))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   458
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   459
        SOME th
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   460
    end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14854
diff changeset
   461
    handle e as THM _ => (message "make_equal: NONE";NONE)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   462
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   463
fun match_consts ignore t (* th *) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   464
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   465
        fun add_consts (Const (c, _), cs) =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36614
diff changeset
   466
            if member (op =) ignore c
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   467
            then cs
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   468
            else insert (op =) c cs
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   469
          | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   470
          | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   471
          | add_consts (_, cs) = cs
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   472
        val t_consts = add_consts(t,[])
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   473
    in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   474
     fn (name,th) =>
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   475
        let
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   476
            val th_consts = add_consts(prop_of th,[])
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   477
        in
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   478
            eq_set (op =) (t_consts, th_consts)
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   479
        end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   480
    end
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   481
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   482
val collect_ignored = fold_rev (fn thm => fn cs =>
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   483
  let
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   484
    val (lhs, rhs) = Logic.dest_equals (prop_of thm);
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   485
    val consts_lhs = Term.add_const_names lhs [];
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   486
    val consts_rhs = Term.add_const_names rhs [];
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   487
    val ignore_lhs = subtract (op =) consts_rhs consts_lhs;
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   488
    val ignore_rhs = subtract (op =) consts_lhs consts_rhs;
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   489
  in
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   490
    fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33038
diff changeset
   491
  end)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   492
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   493
(* set_prop t thms tries to make a theorem with the proposition t from
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   494
one of the theorems thms, by shuffling the propositions around.  If it
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14854
diff changeset
   495
succeeds, SOME theorem is returned, otherwise NONE.  *)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   496
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   497
fun set_prop thy t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   498
    let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43333
diff changeset
   499
        val vars = Misc_Legacy.add_term_frees (t, Misc_Legacy.add_term_vars (t,[]))
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27173
diff changeset
   500
        val closed_t = fold_rev Logic.all vars t
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   501
        val rew_th = norm_term thy closed_t
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22846
diff changeset
   502
        val rhs = Thm.rhs_of rew_th
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   503
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   504
        fun process [] = NONE
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   505
          | process ((name,th)::thms) =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   506
            let
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35408
diff changeset
   507
                val norm_th = Thm.varifyT_global (norm_thm thy (close_thm (Thm.transfer thy th)))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   508
                val triv_th = Thm.trivial rhs
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   509
                val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31244
diff changeset
   510
                val mod_th = case Seq.pull (Thm.bicompose false (*true*) (false,norm_th,0) 1 triv_th) of
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   511
                                 SOME(th,_) => SOME th
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   512
                               | NONE => NONE
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   513
            in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   514
                case mod_th of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   515
                    SOME mod_th =>
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   516
                    let
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   517
                        val closed_th = Thm.equal_elim (Thm.symmetric rew_th) mod_th
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   518
                    in
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   519
                        message ("Shuffler.set_prop succeeded by " ^ name);
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   520
                        SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   521
                    end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   522
                  | NONE => process thms
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   523
            end
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   524
            handle e as THM _ => process thms
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   525
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   526
        fn thms =>
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   527
           case process thms of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   528
               res as SOME (name,th) => if (prop_of th) aconv t
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   529
                                        then res
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   530
                                        else error "Internal error in set_prop"
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   531
             | NONE => NONE
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   532
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   533
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   534
fun find_potential thy t =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   535
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   536
        val shuffles = ShuffleData.get thy
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   537
        val ignored = collect_ignored shuffles []
26662
39483503705f Facts.dest_table;
wenzelm
parents: 26424
diff changeset
   538
        val all_thms =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
   539
          map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (Global_Theory.facts_of thy)))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   540
    in
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33245
diff changeset
   541
        filter (match_consts ignored t) all_thms
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   542
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   543
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
   544
fun gen_shuffle_tac ctxt search thms = SUBGOAL (fn (t, i) =>
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   545
    let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41521
diff changeset
   546
        val thy = Proof_Context.theory_of ctxt
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   547
        val set = set_prop thy t
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   548
        fun process_tac thms st =
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   549
            case set thms of
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   550
                SOME (_,th) => Seq.of_list (compose (th,i,st))
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   551
              | NONE => Seq.empty
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   552
    in
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
   553
        process_tac thms APPEND
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
   554
          (if search then process_tac (find_potential thy t) else no_tac)
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42364
diff changeset
   555
    end)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   556
31244
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   557
fun shuffle_tac ctxt thms =
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   558
  gen_shuffle_tac ctxt false (map (pair "") thms);
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   559
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   560
fun search_tac ctxt =
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   561
  gen_shuffle_tac ctxt true (map (pair "premise") (Assumption.all_prems_of ctxt));
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   562
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   563
fun add_shuffle_rule thm thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   564
    let
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   565
        val shuffles = ShuffleData.get thy
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   566
    in
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   567
        if exists (curry Thm.eq_thm thm) shuffles
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   568
        then (warning ((string_of_thm thm) ^ " already known to the shuffler");
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   569
              thy)
cd0dc678a205 simplified method setup;
wenzelm
parents: 21078
diff changeset
   570
        else ShuffleData.put (thm::shuffles) thy
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   571
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   572
20897
3f8d2834b2c4 attribute: Context.mapping;
wenzelm
parents: 20854
diff changeset
   573
val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I);
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   574
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   575
val setup =
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
   576
  Method.setup @{binding shuffle_tac}
31244
4ed31c673baf fixed superficial ML lapses introduced in b3c7044d47b6;
wenzelm
parents: 31241
diff changeset
   577
    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (shuffle_tac ctxt ths)))
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
   578
    "solve goal by shuffling terms around" #>
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
   579
  Method.setup @{binding search_tac}
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30528
diff changeset
   580
    (Scan.succeed (SIMPLE_METHOD' o search_tac)) "search for suitable theorems" #>
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   581
  add_shuffle_rule weaken #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   582
  add_shuffle_rule equiv_comm #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   583
  add_shuffle_rule imp_comm #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   584
  add_shuffle_rule Drule.norm_hhf_eq #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   585
  add_shuffle_rule Drule.triv_forall_equality #>
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30510
diff changeset
   586
  Attrib.setup @{binding shuffle_rule} (Scan.succeed shuffle_attr) "declare rule for shuffler";
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18127
diff changeset
   587
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   588
end