src/HOL/Import/import.ML
author wenzelm
Sun, 08 Nov 2009 18:43:42 +0100
changeset 33522 737589bb9bb8
parent 32970 fbd2bb2489a8
child 36945 9bec62c10714
permissions -rw-r--r--
adapted Theory_Data; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31244
diff changeset
     1
(*  Title:      HOL/Import/import.ML
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14516
diff changeset
     2
    Author:     Sebastian Skalberg (TU Muenchen)
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14516
diff changeset
     3
*)
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14516
diff changeset
     4
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31244
diff changeset
     5
signature IMPORT =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     6
sig
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32091
diff changeset
     7
    val debug      : bool Unsynchronized.ref
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
     8
    val import_tac : Proof.context -> string * string -> tactic
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 17657
diff changeset
     9
    val setup      : theory -> theory
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    10
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    11
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32970
diff changeset
    12
structure ImportData = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    13
(
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32970
diff changeset
    14
  type T = ProofKernel.thm option array option  (* FIXME mutable array !??!!! *)
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    15
  val empty = NONE
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    16
  val extend = I
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 32970
diff changeset
    17
  fun merge _ = NONE
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22578
diff changeset
    18
)
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14516
diff changeset
    19
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31244
diff changeset
    20
structure Import :> IMPORT =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    21
struct
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    22
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32091
diff changeset
    23
val debug = Unsynchronized.ref false
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    24
fun message s = if !debug then writeln s else ()
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    25
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    26
fun import_tac ctxt (thyname, thmname) =
17370
754b7fcff03e global quick_and_dirty;
wenzelm
parents: 16424
diff changeset
    27
    if ! quick_and_dirty
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32740
diff changeset
    28
    then Skip_Proof.cheat_tac (ProofContext.theory_of ctxt)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    29
    else
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    30
     fn th =>
21590
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    31
        let
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    32
            val thy = ProofContext.theory_of ctxt
21590
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    33
            val prem = hd (prems_of th)
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    34
            val _ = message ("Import_tac: thyname=" ^ thyname ^ ", thmname=" ^ thmname)
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    35
            val _ = message ("Import trying to prove " ^ Syntax.string_of_term ctxt prem)
21590
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    36
            val int_thms = case ImportData.get thy of
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    37
                               NONE => fst (Replay.setup_int_thms thyname thy)
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    38
                             | SOME a => a
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    39
            val proof = snd (ProofKernel.import_proof thyname thmname thy) thy
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    40
            val hol4thm = snd (Replay.replay_proof int_thms thyname thmname proof thy)
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    41
            val thm = snd (ProofKernel.to_isa_thm hol4thm)
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    42
            val rew = ProofKernel.rewrite_hol4_term (concl_of thm) thy
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    43
            val thm = equal_elim rew thm
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    44
            val prew = ProofKernel.rewrite_hol4_term prem thy
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    45
            val prem' = #2 (Logic.dest_equals (prop_of prew))
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
    46
            val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm)
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
    47
            val thm = ProofKernel.disambiguate_frees thm
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
    48
            val _ = message ("Disambiguate: " ^ Display.string_of_thm ctxt thm)
21590
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    49
        in
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    50
            case Shuffler.set_prop thy prem' [("",thm)] of
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    51
                SOME (_,thm) =>
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    52
                let
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    53
                    val _ = if prem' aconv (prop_of thm)
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    54
                            then message "import: Terms match up"
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    55
                            else message "import: Terms DO NOT match up"
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    56
                    val thm' = equal_elim (symmetric prew) thm
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31723
diff changeset
    57
                    val res = Thm.bicompose true (false,thm',0) 1 th
21590
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    58
                in
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    59
                    res
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    60
                end
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    61
              | NONE => (message "import: set_prop didn't succeed"; no_tac th)
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    62
        end
ef7278f553eb tuned spaces/comments;
wenzelm
parents: 18708
diff changeset
    63
31241
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    64
val setup = Method.setup @{binding import}
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    65
  (Scan.lift (Args.name -- Args.name) >>
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    66
    (fn arg => fn ctxt => SIMPLE_METHOD (import_tac ctxt arg)))
b3c7044d47b6 modernized method setup;
wenzelm
parents: 30510
diff changeset
    67
  "import HOL4 theorem"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    68
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    69
end
15707
80b421d8a8be *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
    70