src/HOL/Import/proof_kernel.ML
author wenzelm
Thu, 15 Oct 2009 23:10:35 +0200
changeset 32951 83392f9d303f
parent 32945 63db9da65a19
child 32966 5b21661fe618
permissions -rw-r--r--
space_implode;
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/proof_kernel.ML
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28965
diff changeset
     2
    Author:     Sebastian Skalberg and Steven Obua, TU Muenchen
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14518
diff changeset
     3
*)
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14518
diff changeset
     4
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     5
signature ProofKernel =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     6
sig
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     7
    type hol_type
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     8
    type tag
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
     9
    type term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    10
    type thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    11
    type ('a,'b) subst
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
    12
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    13
    type proof_info
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    14
    datatype proof = Proof of proof_info * proof_content
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    15
         and proof_content
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    16
           = PRefl of term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    17
           | PInstT of proof * (hol_type,hol_type) subst
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    18
           | PSubst of proof list * term * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    19
           | PAbs of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    20
           | PDisch of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    21
           | PMp of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    22
           | PHyp of term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    23
           | PAxm of string * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    24
           | PDef of string * string * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    25
           | PTmSpec of string * string list * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    26
           | PTyDef of string * string * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    27
           | PTyIntro of string * string * string * string * term * term * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    28
           | POracle of tag * term list * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    29
           | PDisk
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    30
           | PSpec of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    31
           | PInst of proof * (term,term) subst
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    32
           | PGen of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    33
           | PGenAbs of proof * term option * term list
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    34
           | PImpAS of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    35
           | PSym of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    36
           | PTrans of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    37
           | PComb of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    38
           | PEqMp of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    39
           | PEqImp of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    40
           | PExists of proof * term * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    41
           | PChoose of term * proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    42
           | PConj of proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    43
           | PConjunct1 of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    44
           | PConjunct2 of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    45
           | PDisj1 of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    46
           | PDisj2 of proof * term
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    47
           | PDisjCases of proof * proof * proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    48
           | PNotI of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    49
           | PNotE of proof
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
    50
           | PContr of proof * term
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    51
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    52
    exception PK of string * string
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    53
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14518
diff changeset
    54
    val get_proof_dir: string -> theory -> string option
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
    55
    val disambiguate_frees : Thm.thm -> Thm.thm
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    56
    val debug : bool Unsynchronized.ref
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    57
    val disk_info_of : proof -> (string * string) option
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    58
    val set_disk_info_of : proof -> string -> string -> unit
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    59
    val mk_proof : proof_content -> proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    60
    val content_of : proof -> proof_content
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    61
    val import_proof : string -> string -> theory -> (theory -> term) option * (theory -> proof)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    62
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    63
    val rewrite_hol4_term: Term.term -> theory -> Thm.thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    64
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    65
    val type_of : term -> hol_type
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    66
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    67
    val get_thm  : string -> string         -> theory -> (theory * thm option)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    68
    val get_def  : string -> string -> term -> theory -> (theory * thm option)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    69
    val get_axiom: string -> string         -> theory -> (theory * thm option)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    70
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    71
    val store_thm : string -> string -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    72
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    73
    val to_isa_thm : thm -> (term * term) list * Thm.thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    74
    val to_isa_term: term -> Term.term
19064
bf19cc5a7899 fixed bugs, added caching
obua
parents: 18929
diff changeset
    75
    val to_hol_thm : Thm.thm -> thm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    76
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    77
    val REFL : term -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    78
    val ASSUME : term -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    79
    val INST_TYPE : (hol_type,hol_type) subst -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    80
    val INST : (term,term)subst -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    81
    val EQ_MP : thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    82
    val EQ_IMP_RULE : thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    83
    val SUBST : thm list -> term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    84
    val DISJ_CASES : thm -> thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    85
    val DISJ1: thm -> term -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    86
    val DISJ2: term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    87
    val IMP_ANTISYM: thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    88
    val SYM : thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    89
    val MP : thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    90
    val GEN : term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    91
    val CHOOSE : term -> thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    92
    val EXISTS : term -> term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    93
    val ABS : term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    94
    val GEN_ABS : term option -> term list -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    95
    val TRANS : thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    96
    val CCONTR : term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    97
    val CONJ : thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    98
    val CONJUNCT1: thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
    99
    val CONJUNCT2: thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   100
    val NOT_INTRO: thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   101
    val NOT_ELIM : thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   102
    val SPEC : term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   103
    val COMB : thm -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   104
    val DISCH: term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   105
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   106
    val type_introduction: string -> string -> string -> string -> string -> term * term -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   107
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   108
    val new_definition : string -> string -> term -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   109
    val new_specification : string -> string -> string list -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   110
    val new_type_definition : string -> string -> string -> thm -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   111
    val new_axiom : string -> term -> theory -> theory * thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   112
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   113
    val prin : term -> unit
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   114
    val protect_factname : string -> string
19067
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   115
    val replay_protect_varname : string -> string -> unit
19068
04b302f2902d cache improvements
obua
parents: 19067
diff changeset
   116
    val replay_add_dump : string -> theory -> theory
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   117
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   118
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   119
structure ProofKernel :> ProofKernel =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   120
struct
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   121
type hol_type = Term.typ
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   122
type term = Term.term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   123
datatype tag = Tag of string list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   124
type ('a,'b) subst = ('a * 'b) list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   125
datatype thm = HOLThm of (Term.term * Term.term) list * Thm.thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   126
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   127
fun hthm2thm (HOLThm (_, th)) = th
17324
0a5eebd5ff58 introduced internal function hthm2thm
obua
parents: 17322
diff changeset
   128
19064
bf19cc5a7899 fixed bugs, added caching
obua
parents: 18929
diff changeset
   129
fun to_hol_thm th = HOLThm ([], th)
17328
7bbfb79eda96 removed clutter
obua
parents: 17324
diff changeset
   130
19068
04b302f2902d cache improvements
obua
parents: 19067
diff changeset
   131
val replay_add_dump = add_dump
04b302f2902d cache improvements
obua
parents: 19067
diff changeset
   132
fun add_dump s thy = (ImportRecorder.add_dump s; replay_add_dump s thy)
04b302f2902d cache improvements
obua
parents: 19067
diff changeset
   133
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   134
datatype proof_info
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   135
  = Info of {disk_info: (string * string) option Unsynchronized.ref}
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   136
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   137
datatype proof = Proof of proof_info * proof_content
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   138
     and proof_content
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   139
       = PRefl of term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   140
       | PInstT of proof * (hol_type,hol_type) subst
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   141
       | PSubst of proof list * term * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   142
       | PAbs of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   143
       | PDisch of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   144
       | PMp of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   145
       | PHyp of term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   146
       | PAxm of string * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   147
       | PDef of string * string * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   148
       | PTmSpec of string * string list * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   149
       | PTyDef of string * string * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   150
       | PTyIntro of string * string * string * string * term * term * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   151
       | POracle of tag * term list * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   152
       | PDisk
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   153
       | PSpec of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   154
       | PInst of proof * (term,term) subst
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   155
       | PGen of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   156
       | PGenAbs of proof * term option * term list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   157
       | PImpAS of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   158
       | PSym of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   159
       | PTrans of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   160
       | PComb of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   161
       | PEqMp of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   162
       | PEqImp of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   163
       | PExists of proof * term * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   164
       | PChoose of term * proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   165
       | PConj of proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   166
       | PConjunct1 of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   167
       | PConjunct2 of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   168
       | PDisj1 of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   169
       | PDisj2 of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   170
       | PDisjCases of proof * proof * proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   171
       | PNotI of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   172
       | PNotE of proof
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   173
       | PContr of proof * term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   174
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   175
exception PK of string * string
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   176
fun ERR f mesg = PK (f,mesg)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   177
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   178
fun print_exn e =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   179
    case e of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   180
        PK (m,s) => (writeln ("PK (" ^ m ^ "): " ^ s); raise e)
17959
8db36a108213 OldGoals;
wenzelm
parents: 17917
diff changeset
   181
      | _ => OldGoals.print_exn e
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   182
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   183
(* Compatibility. *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   184
19264
61e775c03ed8 string_of_mixfix;
wenzelm
parents: 19068
diff changeset
   185
val string_of_mixfix = Pretty.string_of o Syntax.pretty_mixfix;
61e775c03ed8 string_of_mixfix;
wenzelm
parents: 19068
diff changeset
   186
14685
92f34880efe3 more robust output of definitions;
wenzelm
parents: 14673
diff changeset
   187
fun mk_syn thy c =
16427
9975aab75d72 refer to HOL4_PROOFS setting;
wenzelm
parents: 16363
diff changeset
   188
  if Syntax.is_identifier c andalso not (Syntax.is_keyword (Sign.syn_of thy) c) then NoSyn
14685
92f34880efe3 more robust output of definitions;
wenzelm
parents: 14673
diff changeset
   189
  else Syntax.literal c
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   190
14673
3d760a971fde use Syntax.is_identifier;
wenzelm
parents: 14620
diff changeset
   191
fun quotename c =
27353
71c4dd53d4cb moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
wenzelm
parents: 27187
diff changeset
   192
  if Syntax.is_identifier c andalso not (OuterKeyword.is_keyword c) then c else quote c
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   193
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   194
fun simple_smart_string_of_cterm ct =
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   195
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   196
        val thy = Thm.theory_of_cterm ct;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   197
        val {t,T,...} = rep_cterm ct
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   198
        (* Hack to avoid parse errors with Trueprop *)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   199
        val ct  = (cterm_of thy (HOLogic.dest_Trueprop t)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   200
                           handle TERM _ => ct)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   201
    in
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   202
        quote (
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   203
        PrintMode.setmp [] (
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   204
        Library.setmp show_brackets false (
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   205
        Library.setmp show_all_types true (
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   206
        Library.setmp Syntax.ambiguity_is_error false (
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   207
        Library.setmp show_sorts true (Syntax.string_of_term_global thy o Thm.term_of)))))
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   208
        ct)
17652
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   209
    end
b1ef33ebfa17 Release HOL4 and HOLLight Importer.
obua
parents: 17644
diff changeset
   210
19064
bf19cc5a7899 fixed bugs, added caching
obua
parents: 18929
diff changeset
   211
exception SMART_STRING
bf19cc5a7899 fixed bugs, added caching
obua
parents: 18929
diff changeset
   212
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   213
fun smart_string_of_cterm ct =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   214
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   215
        val thy = Thm.theory_of_cterm ct
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   216
        val ctxt = ProofContext.init thy
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   217
        val {t,T,...} = rep_cterm ct
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   218
        (* Hack to avoid parse errors with Trueprop *)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   219
        val ct  = (cterm_of thy (HOLogic.dest_Trueprop t)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   220
                   handle TERM _ => ct)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   221
        fun match u = t aconv u
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   222
        fun G 0 = Library.setmp show_types true (Library.setmp show_sorts true)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   223
          | G 1 = Library.setmp show_brackets true (G 0)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   224
          | G 2 = Library.setmp show_all_types true (G 0)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   225
          | G 3 = Library.setmp show_brackets true (G 2)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   226
          | G _ = raise SMART_STRING
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   227
        fun F n =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   228
            let
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   229
                val str =
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   230
                  Library.setmp show_brackets false (G n (Syntax.string_of_term ctxt o term_of)) ct
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   231
                val u = Syntax.parse_term ctxt str
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   232
                  |> TypeInfer.constrain T |> Syntax.check_term ctxt
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   233
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   234
                if match u
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   235
                then quote str
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   236
                else F (n+1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   237
            end
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   238
            handle ERROR mesg => F (n + 1)
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   239
              | SMART_STRING =>
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   240
                  error ("smart_string failed for: "^ G 0 (Syntax.string_of_term ctxt o term_of) ct)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   241
    in
24634
38db11874724 simplified PrintMode interfaces;
wenzelm
parents: 24630
diff changeset
   242
      PrintMode.setmp [] (Library.setmp Syntax.ambiguity_is_error true F) 0
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   243
    end
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18489
diff changeset
   244
    handle ERROR mesg => simple_smart_string_of_cterm ct
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   245
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   246
val smart_string_of_thm = smart_string_of_cterm o cprop_of
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   247
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
   248
fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th);
32180
37800cb1d378 ML_Context.the_local_context;
wenzelm
parents: 32174
diff changeset
   249
fun prin t = writeln (PrintMode.setmp []
37800cb1d378 ML_Context.the_local_context;
wenzelm
parents: 32174
diff changeset
   250
  (fn () => Syntax.string_of_term (ML_Context.the_local_context ()) t) ());
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   251
fun pth (HOLThm(ren,thm)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   252
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   253
        (*val _ = writeln "Renaming:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   254
        val _ = app (fn(v,w) => (prin v; writeln " -->"; prin w)) ren*)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   255
        val _ = prth thm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   256
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   257
        ()
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   258
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   259
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   260
fun disk_info_of (Proof(Info{disk_info,...},_)) = !disk_info
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   261
fun mk_proof p = Proof(Info{disk_info = Unsynchronized.ref NONE},p)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   262
fun content_of (Proof(_,p)) = p
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   263
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   264
fun set_disk_info_of (Proof(Info{disk_info,...},_)) thyname thmname =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15463
diff changeset
   265
    disk_info := SOME(thyname,thmname)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   266
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   267
structure Lib =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   268
struct
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   269
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   270
fun assoc x =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   271
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   272
        fun F [] = raise PK("Lib.assoc","Not found")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   273
          | F ((x',y)::rest) = if x = x'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   274
                               then y
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   275
                               else F rest
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   276
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   277
        F
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   278
    end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   279
fun i mem L =
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   280
    let fun itr [] = false
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   281
          | itr (a::rst) = i=a orelse itr rst
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   282
    in itr L end;
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   283
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   284
fun [] union S = S
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   285
  | S union [] = S
29289
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
   286
  | (a::rst) union S2 = rst union (insert (op =) a S2)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   287
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   288
fun implode_subst [] = []
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   289
  | implode_subst (x::r::rest) = ((x,r)::(implode_subst rest))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   290
  | implode_subst _ = raise ERR "implode_subst" "malformed substitution list"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   291
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   292
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   293
open Lib
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   294
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   295
structure Tag =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   296
struct
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   297
val empty_tag = Tag []
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   298
fun read name = Tag [name]
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   299
fun merge (Tag tag1) (Tag tag2) = Tag (Lib.union(tag1,tag2))
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   300
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   301
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   302
(* Actual code. *)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   303
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   304
fun get_segment thyname l = (Lib.assoc "s" l
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   305
                             handle PK _ => thyname)
14518
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   306
val get_name : (string * string) list -> string = Lib.assoc "n"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   307
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   308
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   309
    open LazyScan
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   310
    infix 7 |-- --|
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   311
    infix 5 :-- -- ^^
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   312
    infix 3 >>
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   313
    infix 0 ||
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   314
in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   315
exception XML of string
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   316
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   317
datatype xml = Elem of string * (string * string) list * xml list
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   318
datatype XMLtype = XMLty of xml | FullType of hol_type
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   319
datatype XMLterm = XMLtm of xml | FullTerm of term
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   320
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   321
fun pair x y = (x,y)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   322
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   323
fun scan_id toks =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   324
    let
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   325
        val (x,toks2) = one Char.isAlpha toks
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   326
        val (xs,toks3) = any Char.isAlphaNum toks2
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   327
    in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   328
        (String.implode (x::xs),toks3)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   329
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   330
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   331
fun scan_string str c =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   332
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   333
        fun F [] toks = (c,toks)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   334
          | F (c::cs) toks =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   335
            case LazySeq.getItem toks of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   336
                SOME(c',toks') =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   337
                if c = c'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   338
                then F cs toks'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   339
                else raise SyntaxError
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   340
              | NONE => raise SyntaxError
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   341
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   342
        F (String.explode str)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   343
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   344
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   345
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   346
    val scan_entity =
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   347
        (scan_string "amp;" #"&")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   348
            || scan_string "quot;" #"\""
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   349
            || scan_string "gt;" #">"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   350
            || scan_string "lt;" #"<"
14620
1be590fd2422 Minor cleanup of headers and some speedup of the HOL4 import.
skalberg
parents: 14518
diff changeset
   351
            || scan_string "apos;" #"'"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   352
in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   353
fun scan_nonquote toks =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   354
    case LazySeq.getItem toks of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   355
        SOME (c,toks') =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   356
        (case c of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   357
             #"\"" => raise SyntaxError
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   358
           | #"&" => scan_entity toks'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   359
           | c => (c,toks'))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15463
diff changeset
   360
      | NONE => raise SyntaxError
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   361
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   362
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   363
val scan_string = $$ #"\"" |-- repeat scan_nonquote --| $$ #"\"" >>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   364
                     String.implode
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   365
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   366
val scan_attribute = scan_id -- $$ #"=" |-- scan_string
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   367
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   368
val scan_start_of_tag = $$ #"<" |-- scan_id --
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   369
                           repeat ($$ #" " |-- scan_attribute)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   370
14518
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   371
(* The evaluation delay introduced through the 'toks' argument is needed
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   372
for the sake of the SML/NJ (110.9.1) compiler.  Either that or an explicit
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   373
type :-( *)
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   374
fun scan_end_of_tag toks = ($$ #"/" |-- $$ #">" |-- succeed []) toks
c3019a66180f Added a number of explicit type casts and delayed evaluations (all seemingly
skalberg
parents: 14516
diff changeset
   375
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   376
val scan_end_tag = $$ #"<" |-- $$ #"/" |-- scan_id --| $$ #">"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   377
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   378
fun scan_children id = $$ #">" |-- repeat scan_tag -- scan_end_tag >>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   379
                       (fn (chldr,id') => if id = id'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   380
                                          then chldr
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   381
                                          else raise XML "Tag mismatch")
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   382
and scan_tag toks =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   383
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   384
        val ((id,atts),toks2) = scan_start_of_tag toks
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   385
        val (chldr,toks3) = (scan_children id || scan_end_of_tag) toks2
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   386
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   387
        (Elem (id,atts,chldr),toks3)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   388
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   389
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   390
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   391
val type_of = Term.type_of
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   392
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   393
val boolT = Type("bool",[])
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   394
val propT = Type("prop",[])
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   395
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   396
fun mk_defeq name rhs thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   397
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   398
        val ty = type_of rhs
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   399
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   400
        Logic.mk_equals (Const(Sign.intern_const thy name,ty),rhs)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   401
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   402
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   403
fun mk_teq name rhs thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   404
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   405
        val ty = type_of rhs
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   406
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   407
        HOLogic.mk_eq (Const(Sign.intern_const thy name,ty),rhs)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   408
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   409
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   410
fun intern_const_name thyname const thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   411
    case get_hol4_const_mapping thyname const thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   412
        SOME (_,cname,_) => cname
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15463
diff changeset
   413
      | NONE => (case get_hol4_const_renaming thyname const thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   414
                     SOME cname => Sign.intern_const thy (thyname ^ "." ^ cname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   415
                   | NONE => Sign.intern_const thy (thyname ^ "." ^ const))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   416
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   417
fun intern_type_name thyname const thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   418
    case get_hol4_type_mapping thyname const thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   419
        SOME (_,cname) => cname
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   420
      | NONE => Sign.intern_const thy (thyname ^ "." ^ const)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   421
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   422
fun mk_vartype name = TFree(name,["HOL.type"])
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   423
fun mk_thy_type thy Thy Tyop Args = Type(intern_type_name Thy Tyop thy,Args)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   424
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   425
val mk_var = Free
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   426
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   427
fun dom_rng (Type("fun",[dom,rng])) = (dom,rng)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   428
  | dom_rng _ = raise ERR "dom_rng" "Not a functional type"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   429
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16427
diff changeset
   430
fun mk_thy_const thy Thy Nam Ty = Const(intern_const_name Thy Nam thy,Ty)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   431
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   432
local
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   433
  fun get_const sg thyname name =
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   434
    (case Sign.const_type sg name of
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   435
      SOME ty => Const (name, ty)
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   436
    | NONE => raise ERR "get_type" (name ^ ": No such constant"))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   437
in
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16427
diff changeset
   438
fun prim_mk_const thy Thy Nam =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   439
    let
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   440
      val name = intern_const_name Thy Nam thy
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   441
      val cmaps = HOL4ConstMaps.get thy
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   442
    in
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   443
      case StringPair.lookup cmaps (Thy,Nam) of
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   444
        SOME(_,_,SOME ty) => Const(name,ty)
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
   445
      | _ => get_const thy Thy name
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   446
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   447
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   448
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   449
fun mk_comb(f,a) = f $ a
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   450
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   451
(* Needed for HOL Light *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   452
fun protect_tyvarname s =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   453
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   454
        fun no_quest s =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   455
            if Char.contains s #"?"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   456
            then String.translate (fn #"?" => "q_" | c => Char.toString c) s
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   457
            else s
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   458
        fun beg_prime s =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   459
            if String.isPrefix "'" s
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   460
            then s
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   461
            else "'" ^ s
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   462
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   463
        s |> no_quest |> beg_prime
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   464
    end
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   465
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   466
val protected_varnames = Unsynchronized.ref (Symtab.empty:string Symtab.table)
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   467
val invented_isavar = Unsynchronized.ref 0
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   468
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   469
fun innocent_varname s = Syntax.is_identifier s andalso not (String.isPrefix "u_" s)
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   470
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18489
diff changeset
   471
val check_name_thy = theory "Main"
17592
ece268908438 add debug messages
obua
parents: 17490
diff changeset
   472
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18489
diff changeset
   473
fun valid_boundvarname s =
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   474
  can (fn () => Syntax.read_term_global check_name_thy ("SOME "^s^". True")) ();
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18489
diff changeset
   475
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18489
diff changeset
   476
fun valid_varname s =
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   477
  can (fn () => Syntax.read_term_global check_name_thy s) ();
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   478
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   479
fun protect_varname s =
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   480
    if innocent_varname s andalso valid_varname s then s else
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   481
    case Symtab.lookup (!protected_varnames) s of
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   482
      SOME t => t
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   483
    | NONE =>
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   484
      let
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   485
          val _ = Unsynchronized.inc invented_isavar
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   486
          val t = "u_" ^ string_of_int (!invented_isavar)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   487
          val _ = ImportRecorder.protect_varname s t
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   488
          val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   489
      in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   490
          t
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   491
      end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   492
19067
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   493
exception REPLAY_PROTECT_VARNAME of string*string*string
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   494
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   495
fun replay_protect_varname s t =
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   496
        case Symtab.lookup (!protected_varnames) s of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   497
          SOME t' => raise REPLAY_PROTECT_VARNAME (s, t, t')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   498
        | NONE =>
19067
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   499
          let
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
   500
              val _ = Unsynchronized.inc invented_isavar
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   501
              val t = "u_" ^ string_of_int (!invented_isavar)
19067
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   502
              val _ = protected_varnames := Symtab.update (s, t) (!protected_varnames)
c0321d7d6b3d variable counter is now also cached
obua
parents: 19066
diff changeset
   503
          in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   504
              ()
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   505
          end
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   506
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   507
fun protect_boundvarname s = if innocent_varname s andalso valid_boundvarname s then s else "u"
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   508
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   509
fun mk_lambda (v as Free (x, T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   510
  | mk_lambda (v as Var ((x, _), T)) t = Abs (protect_boundvarname x, T, abstract_over (v, t))
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   511
  | mk_lambda v t = raise TERM ("lambda", [v, t]);
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   512
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   513
fun replacestr x y s =
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   514
let
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   515
  val xl = explode x
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   516
  val yl = explode y
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   517
  fun isprefix [] ys = true
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   518
    | isprefix (x::xs) (y::ys) = if x = y then isprefix xs ys else false
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   519
    | isprefix _ _ = false
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   520
  fun isp s = isprefix xl s
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   521
  fun chg s = yl@(List.drop (s, List.length xl))
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   522
  fun r [] = []
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   523
    | r (S as (s::ss)) = if isp S then r (chg S) else s::(r ss)
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   524
in
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   525
  implode(r (explode s))
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   526
end
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   527
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   528
fun protect_factname s = replacestr "." "_dot_" s
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   529
fun unprotect_factname s = replacestr "_dot_" "." s
17440
df77edc4f5d0 fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
obua
parents: 17412
diff changeset
   530
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   531
val ty_num_prefix = "N_"
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   532
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   533
fun startsWithDigit s = Char.isDigit (hd (String.explode s))
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   534
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   535
fun protect_tyname tyn =
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   536
  let
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   537
    val tyn' =
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   538
      if String.isPrefix ty_num_prefix tyn then raise (ERR "protect_ty_name" ("type name '"^tyn^"' is reserved")) else
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   539
      (if startsWithDigit tyn then ty_num_prefix^tyn else tyn)
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   540
  in
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   541
    tyn'
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   542
  end
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   543
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   544
fun protect_constname tcn = tcn
17444
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   545
 (* if tcn = ".." then "dotdot"
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   546
  else if tcn = "==" then "eqeq"
a389e5892691 1) mapped .. and == constants
obua
parents: 17440
diff changeset
   547
  else tcn*)
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   548
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   549
structure TypeNet =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   550
struct
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   551
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   552
fun get_type_from_index thy thyname types is =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   553
    case Int.fromString is of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   554
        SOME i => (case Array.sub(types,i) of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   555
                       FullType ty => ty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   556
                     | XMLty xty =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   557
                       let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   558
                           val ty = get_type_from_xml thy thyname types xty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   559
                           val _  = Array.update(types,i,FullType ty)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   560
                       in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   561
                           ty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   562
                       end)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   563
      | NONE => raise ERR "get_type_from_index" "Bad index"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   564
and get_type_from_xml thy thyname types =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   565
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   566
        fun gtfx (Elem("tyi",[("i",iS)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   567
                 get_type_from_index thy thyname types iS
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   568
          | gtfx (Elem("tyc",atts,[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   569
            mk_thy_type thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   570
                        (get_segment thyname atts)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   571
                        (protect_tyname (get_name atts))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   572
                        []
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   573
          | gtfx (Elem("tyv",[("n",s)],[])) = mk_vartype (protect_tyvarname s)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   574
          | gtfx (Elem("tya",[],(Elem("tyc",atts,[]))::tys)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   575
            mk_thy_type thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   576
                        (get_segment thyname atts)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   577
                        (protect_tyname (get_name atts))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   578
                        (map gtfx tys)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   579
          | gtfx _ = raise ERR "get_type" "Bad type"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   580
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   581
        gtfx
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   582
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   583
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   584
fun input_types thyname (Elem("tylist",[("i",i)],xtys)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   585
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   586
        val types = Array.array(valOf (Int.fromString i),XMLty (Elem("",[],[])))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   587
        fun IT _ [] = ()
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   588
          | IT n (xty::xtys) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   589
            (Array.update(types,n,XMLty xty);
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   590
             IT (n+1) xtys)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   591
        val _ = IT 0 xtys
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   592
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   593
        types
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   594
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   595
  | input_types _ _ = raise ERR "input_types" "Bad type list"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   596
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   597
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   598
structure TermNet =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   599
struct
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   600
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   601
fun get_term_from_index thy thyname types terms is =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   602
    case Int.fromString is of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   603
        SOME i => (case Array.sub(terms,i) of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   604
                       FullTerm tm => tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   605
                     | XMLtm xtm =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   606
                       let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   607
                           val tm = get_term_from_xml thy thyname types terms xtm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   608
                           val _  = Array.update(terms,i,FullTerm tm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   609
                       in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   610
                           tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   611
                       end)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   612
      | NONE => raise ERR "get_term_from_index" "Bad index"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   613
and get_term_from_xml thy thyname types terms =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   614
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   615
        fun get_type [] = NONE
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   616
          | get_type [ty] = SOME (TypeNet.get_type_from_xml thy thyname types ty)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   617
          | get_type _ = raise ERR "get_term" "Bad type"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   618
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   619
        fun gtfx (Elem("tmv",[("n",name),("t",tyi)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   620
            mk_var(protect_varname name,TypeNet.get_type_from_index thy thyname types tyi)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   621
          | gtfx (Elem("tmc",atts,[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   622
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   623
                val segment = get_segment thyname atts
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   624
                val name = protect_constname(get_name atts)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   625
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   626
                mk_thy_const thy segment name (TypeNet.get_type_from_index thy thyname types (Lib.assoc "t" atts))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   627
                handle PK _ => prim_mk_const thy segment name
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   628
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   629
          | gtfx (Elem("tma",[("f",tmf),("a",tma)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   630
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   631
                val f = get_term_from_index thy thyname types terms tmf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   632
                val a = get_term_from_index thy thyname types terms tma
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   633
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   634
                mk_comb(f,a)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   635
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   636
          | gtfx (Elem("tml",[("x",tmx),("a",tma)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   637
            let
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   638
                val x = get_term_from_index thy thyname types terms tmx
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
   639
                val a = get_term_from_index thy thyname types terms tma
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   640
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   641
                mk_lambda x a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   642
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   643
          | gtfx (Elem("tmi",[("i",iS)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   644
            get_term_from_index thy thyname types terms iS
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   645
          | gtfx (Elem(tag,_,_)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   646
            raise ERR "get_term" ("Not a term: "^tag)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   647
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   648
        gtfx
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   649
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   650
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   651
fun input_terms thyname types (Elem("tmlist",[("i",i)],xtms)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   652
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   653
        val terms = Array.array(valOf(Int.fromString i),XMLtm (Elem("",[],[])))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   654
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   655
        fun IT _ [] = ()
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   656
          | IT n (xtm::xtms) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   657
            (Array.update(terms,n,XMLtm xtm);
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   658
             IT (n+1) xtms)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   659
        val _ = IT 0 xtms
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   660
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   661
        terms
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   662
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   663
  | input_terms _ _ _ = raise ERR "input_terms" "Bad term list"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   664
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   665
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   666
fun get_proof_dir (thyname:string) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   667
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   668
        val import_segment =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   669
            case get_segment2 thyname thy of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   670
                SOME seg => seg
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   671
              | NONE => get_import_segment thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   672
        val path = space_explode ":" (getenv "HOL4_PROOFS")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   673
        fun find [] = NONE
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   674
          | find (p::ps) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   675
            (let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   676
                 val dir = OS.Path.joinDirFile {dir = p,file=import_segment}
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   677
             in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   678
                 if OS.FileSys.isDir dir
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   679
                 then SOME dir
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   680
                 else find ps
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   681
             end) handle OS.SysErr _ => find ps
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   682
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   683
        Option.map (fn p => OS.Path.joinDirFile {dir = p, file = thyname}) (find path)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   684
    end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   685
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   686
fun proof_file_name thyname thmname thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   687
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   688
        val path = case get_proof_dir thyname thy of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   689
                       SOME p => p
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   690
                     | NONE => error "Cannot find proof files"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   691
        val _ = OS.FileSys.mkDir path handle OS.SysErr _ => ()
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   692
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   693
        OS.Path.joinDirFile {dir = path, file = OS.Path.joinBaseExt {base = (unprotect_factname thmname), ext = SOME "prf"}}
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   694
    end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   695
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   696
fun xml_to_proof thyname types terms prf thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   697
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   698
        val xml_to_hol_type = TypeNet.get_type_from_xml thy thyname types
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   699
        val xml_to_term = TermNet.get_term_from_xml thy thyname types terms
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   700
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   701
        fun index_to_term is =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   702
            TermNet.get_term_from_index thy thyname types terms is
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   703
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   704
        fun x2p (Elem("prefl",[("i",is)],[])) = mk_proof (PRefl (index_to_term is))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   705
          | x2p (Elem("pinstt",[],p::lambda)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   706
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   707
                val p = x2p p
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   708
                val lambda = implode_subst (map xml_to_hol_type lambda)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   709
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   710
                mk_proof (PInstT(p,lambda))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   711
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   712
          | x2p (Elem("psubst",[("i",is)],prf::prfs)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   713
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   714
                val tm = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   715
                val prf = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   716
                val prfs = map x2p prfs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   717
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   718
                mk_proof (PSubst(prfs,tm,prf))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   719
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   720
          | x2p (Elem("pabs",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   721
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   722
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   723
                val t = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   724
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   725
                mk_proof (PAbs (p,t))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   726
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   727
          | x2p (Elem("pdisch",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   728
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   729
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   730
                val t = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   731
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   732
                mk_proof (PDisch (p,t))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   733
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   734
          | x2p (Elem("pmp",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   735
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   736
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   737
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   738
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   739
                mk_proof (PMp(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   740
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   741
          | x2p (Elem("phyp",[("i",is)],[])) = mk_proof (PHyp (index_to_term is))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   742
          | x2p (Elem("paxiom",[("n",n),("i",is)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   743
            mk_proof (PAxm(n,index_to_term is))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   744
          | x2p (Elem("pfact",atts,[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   745
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   746
                val thyname = get_segment thyname atts
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   747
                val thmname = protect_factname (get_name atts)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   748
                val p = mk_proof PDisk
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   749
                val _  = set_disk_info_of p thyname thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   750
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   751
                p
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   752
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   753
          | x2p (Elem("pdef",[("s",seg),("n",name),("i",is)],[])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   754
            mk_proof (PDef(seg,protect_constname name,index_to_term is))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   755
          | x2p (Elem("ptmspec",[("s",seg)],p::names)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   756
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   757
                val names = map (fn Elem("name",[("n",name)],[]) => name
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   758
                                  | _ => raise ERR "x2p" "Bad proof (ptmspec)") names
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   759
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   760
                mk_proof (PTmSpec(seg,names,x2p p))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   761
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   762
          | x2p (Elem("ptyintro",[("s",seg),("n",name),("a",abs_name),("r",rep_name)],[xP,xt,p])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   763
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   764
                val P = xml_to_term xP
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   765
                val t = xml_to_term xt
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   766
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   767
                mk_proof (PTyIntro(seg,protect_tyname name,protect_constname abs_name,protect_constname rep_name,P,t,x2p p))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   768
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   769
          | x2p (Elem("ptydef",[("s",seg),("n",name)],[p])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   770
            mk_proof (PTyDef(seg,protect_tyname name,x2p p))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   771
          | x2p (xml as Elem("poracle",[],chldr)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   772
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   773
                val (oracles,terms) = List.partition (fn (Elem("oracle",_,_)) => true | _ => false) chldr
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   774
                val ors = map (fn (Elem("oracle",[("n",name)],[])) => name | xml => raise ERR "x2p" "bad oracle") oracles
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   775
                val (c,asl) = case terms of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   776
                                  [] => raise ERR "x2p" "Bad oracle description"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   777
                                | (hd::tl) => (hd,tl)
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
   778
                val tg = List.foldr (fn (oracle,tg) => Tag.merge (Tag.read oracle) tg) Tag.empty_tag ors
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   779
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   780
                mk_proof (POracle(tg,map xml_to_term asl,xml_to_term c))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   781
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   782
          | x2p (Elem("pspec",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   783
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   784
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   785
                val tm = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   786
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   787
                mk_proof (PSpec(p,tm))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   788
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   789
          | x2p (Elem("pinst",[],p::theta)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   790
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   791
                val p = x2p p
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   792
                val theta = implode_subst (map xml_to_term theta)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   793
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   794
                mk_proof (PInst(p,theta))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   795
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   796
          | x2p (Elem("pgen",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   797
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   798
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   799
                val tm = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   800
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   801
                mk_proof (PGen(p,tm))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   802
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   803
          | x2p (Elem("pgenabs",[],prf::tms)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   804
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   805
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   806
                val tml = map xml_to_term tms
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   807
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   808
                mk_proof (PGenAbs(p,NONE,tml))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   809
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   810
          | x2p (Elem("pgenabs",[("i",is)],prf::tms)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   811
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   812
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   813
                val tml = map xml_to_term tms
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   814
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   815
                mk_proof (PGenAbs(p,SOME (index_to_term is),tml))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   816
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   817
          | x2p (Elem("pimpas",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   818
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   819
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   820
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   821
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   822
                mk_proof (PImpAS(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   823
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   824
          | x2p (Elem("psym",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   825
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   826
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   827
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   828
                mk_proof (PSym p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   829
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   830
          | x2p (Elem("ptrans",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   831
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   832
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   833
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   834
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   835
                mk_proof (PTrans(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   836
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   837
          | x2p (Elem("pcomb",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   838
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   839
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   840
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   841
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   842
                mk_proof (PComb(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   843
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   844
          | x2p (Elem("peqmp",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   845
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   846
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   847
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   848
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   849
                mk_proof (PEqMp(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   850
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   851
          | x2p (Elem("peqimp",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   852
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   853
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   854
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   855
                mk_proof (PEqImp p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   856
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   857
          | x2p (Elem("pexists",[("e",ise),("w",isw)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   858
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   859
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   860
                val ex = index_to_term ise
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   861
                val w = index_to_term isw
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   862
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   863
                mk_proof (PExists(p,ex,w))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   864
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   865
          | x2p (Elem("pchoose",[("i",is)],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   866
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   867
                val v  = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   868
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   869
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   870
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   871
                mk_proof (PChoose(v,p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   872
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   873
          | x2p (Elem("pconj",[],[prf1,prf2])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   874
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   875
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   876
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   877
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   878
                mk_proof (PConj(p1,p2))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   879
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   880
          | x2p (Elem("pconjunct1",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   881
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   882
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   883
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   884
                mk_proof (PConjunct1 p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   885
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   886
          | x2p (Elem("pconjunct2",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   887
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   888
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   889
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   890
                mk_proof (PConjunct2 p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   891
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   892
          | x2p (Elem("pdisj1",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   893
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   894
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   895
                val t = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   896
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   897
                mk_proof (PDisj1 (p,t))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   898
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   899
          | x2p (Elem("pdisj2",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   900
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   901
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   902
                val t = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   903
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   904
                mk_proof (PDisj2 (p,t))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   905
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   906
          | x2p (Elem("pdisjcases",[],[prf1,prf2,prf3])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   907
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   908
                val p1 = x2p prf1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   909
                val p2 = x2p prf2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   910
                val p3 = x2p prf3
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   911
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   912
                mk_proof (PDisjCases(p1,p2,p3))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   913
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   914
          | x2p (Elem("pnoti",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   915
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   916
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   917
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   918
                mk_proof (PNotI p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   919
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   920
          | x2p (Elem("pnote",[],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   921
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   922
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   923
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   924
                mk_proof (PNotE p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   925
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   926
          | x2p (Elem("pcontr",[("i",is)],[prf])) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   927
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   928
                val p = x2p prf
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   929
                val t = index_to_term is
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   930
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   931
                mk_proof (PContr (p,t))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   932
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   933
          | x2p xml = raise ERR "x2p" "Bad proof"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   934
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   935
        x2p prf
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   936
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   937
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   938
fun import_proof_concl thyname thmname thy =
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   939
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   940
        val is = TextIO.openIn(proof_file_name thyname thmname thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   941
        val (proof_xml,_) = scan_tag (LazySeq.of_instream is)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   942
        val _ = TextIO.closeIn is
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   943
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   944
        case proof_xml of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   945
            Elem("proof",[],xtypes::xterms::prf::rest) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   946
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   947
                val types = TypeNet.input_types thyname xtypes
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   948
                val terms = TermNet.input_terms thyname types xterms
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   949
                fun f xtm thy = TermNet.get_term_from_xml thy thyname types terms xtm
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   950
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   951
                case rest of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   952
                    [] => NONE
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   953
                  | [xtm] => SOME (f xtm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   954
                  | _ => raise ERR "import_proof" "Bad argument list"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   955
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   956
          | _ => raise ERR "import_proof" "Bad proof"
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   957
    end
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
   958
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   959
fun import_proof thyname thmname thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   960
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   961
        val is = TextIO.openIn(proof_file_name thyname thmname thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   962
        val (proof_xml,_) = scan_tag (LazySeq.of_instream is)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   963
        val _ = TextIO.closeIn is
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
   964
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   965
        case proof_xml of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   966
            Elem("proof",[],xtypes::xterms::prf::rest) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   967
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   968
                val types = TypeNet.input_types thyname xtypes
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   969
                val terms = TermNet.input_terms thyname types xterms
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   970
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   971
                (case rest of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   972
                     [] => NONE
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   973
                   | [xtm] => SOME (fn thy => TermNet.get_term_from_xml thy thyname types terms xtm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   974
                   | _ => raise ERR "import_proof" "Bad argument list",
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   975
                 xml_to_proof thyname types terms prf)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   976
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   977
          | _ => raise ERR "import_proof" "Bad proof"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   978
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   979
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   980
fun uniq_compose m th i st =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   981
    let
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31723
diff changeset
   982
        val res = Thm.bicompose false (false,th,m) i st
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   983
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   984
        case Seq.pull res of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   985
            SOME (th,rest) => (case Seq.pull rest of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   986
                                   SOME _ => raise ERR "uniq_compose" "Not unique!"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   987
                                 | NONE => th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
   988
          | NONE => raise ERR "uniq_compose" "No result"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   989
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   990
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   991
val reflexivity_thm = thm "refl"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   992
val substitution_thm = thm "subst"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   993
val mp_thm = thm "mp"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   994
val imp_antisym_thm = thm "light_imp_as"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   995
val disch_thm = thm "impI"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   996
val ccontr_thm = thm "ccontr"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   997
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   998
val meta_eq_to_obj_eq_thm = thm "meta_eq_to_obj_eq"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
   999
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1000
val gen_thm = thm "HOLallI"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1001
val choose_thm = thm "exE"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1002
val exists_thm = thm "exI"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1003
val conj_thm = thm "conjI"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1004
val conjunct1_thm = thm "conjunct1"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1005
val conjunct2_thm = thm "conjunct2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1006
val spec_thm = thm "spec"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1007
val disj_cases_thm = thm "disjE"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1008
val disj1_thm = thm "disjI1"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1009
val disj2_thm = thm "disjI2"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1010
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1011
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1012
    val th = thm "not_def"
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1013
    val thy = theory_of_thm th
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1014
    val pp = reflexive (cterm_of thy (Const("Trueprop",boolT-->propT)))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1015
in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1016
val not_elim_thm = combination pp th
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1017
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1018
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1019
val not_intro_thm = symmetric not_elim_thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1020
val abs_thm = thm "ext"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1021
val trans_thm = thm "trans"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1022
val symmetry_thm = thm "sym"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1023
val transitivity_thm = thm "trans"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1024
val eqmp_thm = thm "iffD1"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1025
val eqimp_thm = thm "HOL4Setup.eq_imp"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1026
val comb_thm = thm "cong"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1027
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1028
(* Beta-eta normalizes a theorem (only the conclusion, not the *
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1029
hypotheses!)  *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1030
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1031
fun beta_eta_thm th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1032
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1033
        val th1 = Thm.equal_elim (Thm.beta_conversion true (cprop_of th))  th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1034
        val th2 = Thm.equal_elim (Thm.eta_conversion       (cprop_of th1)) th1
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1035
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1036
        th2
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1037
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1038
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1039
fun implies_elim_all th =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
  1040
    Library.foldl (fn (th,p) => implies_elim th (assume p)) (th,cprems_of th)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1041
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1042
fun norm_hyps th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1043
    th |> beta_eta_thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1044
       |> implies_elim_all
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1045
       |> implies_intr_hyps
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1046
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1047
fun mk_GEN v th sg =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1048
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1049
        val c = HOLogic.dest_Trueprop (concl_of th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1050
        val cv = cterm_of sg v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1051
        val lc = Term.lambda v c
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1052
        val clc = Thm.cterm_of sg lc
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1053
        val cvty = ctyp_of_term cv
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1054
        val th1 = implies_elim_all th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1055
        val th2 = beta_eta_thm (forall_intr cv th1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1056
        val th3 = th2 COMP (beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME clc] gen_thm))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1057
        val c = prop_of th3
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1058
        val vname = fst(dest_Free v)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1059
        val (cold,cnew) = case c of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1060
                              tpc $ (Const("All",allT) $ Abs(oldname,ty,body)) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1061
                              (Abs(oldname,dummyT,Bound 0),Abs(vname,dummyT,Bound 0))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1062
                            | tpc $ (Const("All",allT) $ rest) => (tpc,tpc)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1063
                            | _ => raise ERR "mk_GEN" "Unknown conclusion"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1064
        val th4 = Thm.rename_boundvars cold cnew th3
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1065
        val res = implies_intr_hyps th4
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1066
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1067
        res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1068
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1069
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1070
fun rearrange sg tm th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1071
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1072
        val tm' = Envir.beta_eta_contract tm
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31723
diff changeset
  1073
        fun find []      n = Thm.permute_prems 0 1 (implies_intr (Thm.cterm_of sg tm) th)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1074
          | find (p::ps) n = if tm' aconv (Envir.beta_eta_contract p)
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31723
diff changeset
  1075
                             then Thm.permute_prems n 1 th
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1076
                             else find ps (n+1)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1077
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1078
        find (prems_of th) 0
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1079
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1080
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1081
fun zip (x::xs) (y::ys) = (x,y)::(zip xs ys)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1082
  | zip [] [] = []
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1083
  | zip _ _ = raise ERR "zip" "arguments not of same length"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1084
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1085
fun mk_INST dom rng th =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1086
    th |> forall_intr_list dom
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1087
       |> forall_elim_list rng
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1088
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1089
val collect_vars =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1090
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1091
        fun F vars (Bound _) = vars
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1092
          | F vars (tm as Free _) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1093
            if tm mem vars
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1094
            then vars
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1095
            else (tm::vars)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1096
          | F vars (Const _) = vars
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1097
          | F vars (tm1 $ tm2) = F (F vars tm1) tm2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1098
          | F vars (Abs(_,_,body)) = F vars body
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1099
          | F vars (Var _) = raise ERR "collect_vars" "Schematic variable found"
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1100
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1101
        F []
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1102
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1103
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1104
(* Code for disambiguating variablenames (wrt. types) *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1105
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1106
val disamb_info_empty = {vars=[],rens=[]}
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1107
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1108
fun rens_of {vars,rens} = rens
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1109
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1110
fun name_of_var (Free(vname,_)) = vname
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1111
  | name_of_var _ = raise ERR "name_of_var" "Not a variable"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1112
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1113
fun disamb_term_from info tm = (info, tm)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1114
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1115
fun swap (x,y) = (y,x)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1116
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1117
fun has_ren (HOLThm _) = false
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1118
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1119
fun prinfo {vars,rens} = (writeln "Vars:";
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1120
                          app prin vars;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1121
                          writeln "Renaming:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1122
                          app (fn(x,y)=>(prin x; writeln " -->"; prin y)) rens)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1123
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1124
fun disamb_thm_from info (HOLThm (_,thm)) = (info, thm)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1125
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1126
fun disamb_terms_from info tms = (info, tms)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1127
17324
0a5eebd5ff58 introduced internal function hthm2thm
obua
parents: 17322
diff changeset
  1128
fun disamb_thms_from info hthms = (info, map hthm2thm hthms)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1129
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1130
fun disamb_term tm   = disamb_term_from disamb_info_empty tm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1131
fun disamb_terms tms = disamb_terms_from disamb_info_empty tms
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1132
fun disamb_thm thm   = disamb_thm_from disamb_info_empty thm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1133
fun disamb_thms thms = disamb_thms_from disamb_info_empty thms
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1134
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1135
fun norm_hthm sg (hth as HOLThm _) = hth
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1136
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1137
(* End of disambiguating code *)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1138
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1139
fun disambiguate_frees thm =
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1140
    let
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1141
      fun ERR s = error ("Drule.disambiguate_frees: "^s)
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1142
      val ct = cprop_of thm
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1143
      val t = term_of ct
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1144
      val thy = theory_of_cterm ct
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28965
diff changeset
  1145
      val frees = OldTerm.term_frees t
29281
b22ccb3998db eliminated OldTerm.add_term_free_names;
wenzelm
parents: 29270
diff changeset
  1146
      val freenames = Term.add_free_names t []
b22ccb3998db eliminated OldTerm.add_term_free_names;
wenzelm
parents: 29270
diff changeset
  1147
      val is_old_name = member (op =) freenames
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1148
      fun name_of (Free (n, _)) = n
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1149
        | name_of _ = ERR "name_of"
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1150
      fun new_name' bump map n =
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1151
          let val n' = n^bump in
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1152
            if is_old_name n' orelse Symtab.lookup map n' <> NONE then
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1153
              new_name' (Symbol.bump_string bump) map n
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1154
            else
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1155
              n'
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1156
          end
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1157
      val new_name = new_name' "a"
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1158
      fun replace_name n' (Free (n, t)) = Free (n', t)
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1159
        | replace_name n' _ = ERR "replace_name"
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1160
      (* map: old or fresh name -> old free,
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1161
         invmap: old free which has fresh name assigned to it -> fresh name *)
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1162
      fun dis (v, mapping as (map,invmap)) =
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1163
          let val n = name_of v in
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1164
            case Symtab.lookup map n of
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1165
              NONE => (Symtab.update (n, v) map, invmap)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1166
            | SOME v' =>
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1167
              if v=v' then
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1168
                mapping
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1169
              else
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1170
                let val n' = new_name map n in
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1171
                  (Symtab.update (n', v) map,
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1172
                   Termtab.update (v, n') invmap)
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1173
                end
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1174
          end
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1175
    in
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1176
      if (length freenames = length frees) then
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1177
        thm
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1178
      else
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1179
        let
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1180
          val (_, invmap) =
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1181
              List.foldl dis (Symtab.empty, Termtab.empty) frees
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1182
          fun make_subst ((oldfree, newname), (intros, elims)) =
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1183
              (cterm_of thy oldfree :: intros,
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1184
               cterm_of thy (replace_name newname oldfree) :: elims)
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1185
          val (intros, elims) = List.foldl make_subst ([], []) (Termtab.dest invmap)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1186
        in
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1187
          forall_elim_list elims (forall_intr_list intros thm)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1188
        end
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1189
    end
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1190
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
  1191
val debug = Unsynchronized.ref false
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1192
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1193
fun if_debug f x = if !debug then f x else ()
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1194
val message = if_debug writeln
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1195
31945
d5f186aa0bed structure Thm: less pervasive names;
wenzelm
parents: 31723
diff changeset
  1196
val conjE_helper = Thm.permute_prems 0 1 conjE
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1197
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1198
fun get_hol4_thm thyname thmname thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1199
    case get_hol4_theorem thyname thmname thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1200
        SOME hth => SOME (HOLThm hth)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1201
      | NONE =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1202
        let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1203
            val pending = HOL4Pending.get thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1204
        in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1205
            case StringPair.lookup pending (thyname,thmname) of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1206
                SOME hth => SOME (HOLThm hth)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1207
              | NONE => NONE
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1208
        end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1209
29289
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1210
fun non_trivial_term_consts t = fold_aterms
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1211
  (fn Const (c, _) =>
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1212
      if c = "Trueprop" orelse c = "All" orelse c = "op -->" orelse c = "op &" orelse c = "op ="
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1213
      then I else insert (op =) c
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1214
    | _ => I) t [];
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1215
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1216
fun match_consts t (* th *) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1217
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1218
        fun add_consts (Const (c, _), cs) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1219
            (case c of
29289
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1220
                 "op =" => insert (op =) "==" cs
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1221
               | "op -->" => insert (op =) "==>" cs
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1222
               | "All" => cs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1223
               | "all" => cs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1224
               | "op &" => cs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1225
               | "Trueprop" => cs
29289
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1226
               | _ => insert (op =) c cs)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1227
          | add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1228
          | add_consts (Abs (_, _, t), cs) = add_consts (t, cs)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1229
          | add_consts (_, cs) = cs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1230
        val t_consts = add_consts(t,[])
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1231
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1232
        fn th => eq_set(t_consts,add_consts(prop_of th,[]))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1233
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1234
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1235
fun split_name str =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1236
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1237
        val sub = Substring.full str
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1238
        val (f,idx) = apsnd Substring.string (Substring.splitr Char.isDigit sub)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1239
        val (newstr,u) = pairself Substring.string (Substring.splitr (fn c => c = #"_") f)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1240
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1241
        if not (idx = "") andalso u = "_"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1242
        then SOME (newstr,valOf(Int.fromString idx))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1243
        else NONE
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1244
    end
28397
389c5e494605 handle _ should be avoided (spurious Interrupt will spoil the game);
wenzelm
parents: 27691
diff changeset
  1245
    handle _ => NONE  (* FIXME avoid handle _ *)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1246
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1247
fun rewrite_hol4_term t thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1248
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1249
        val hol4rews1 = map (Thm.transfer thy) (HOL4Rewrites.get thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1250
        val hol4ss = Simplifier.theory_context thy empty_ss
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1251
          setmksimps single addsimps hol4rews1
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1252
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1253
        Thm.transfer thy (Simplifier.full_rewrite hol4ss (cterm_of thy t))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1254
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1255
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1256
fun get_isabelle_thm thyname thmname hol4conc thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1257
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1258
        val (info,hol4conc') = disamb_term hol4conc
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1259
        val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1260
        val isaconc =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1261
            case concl_of i2h_conc of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1262
                Const("==",_) $ lhs $ _ => lhs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1263
              | _ => error "get_isabelle_thm" "Bad rewrite rule"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1264
        val _ = (message "Original conclusion:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1265
                 if_debug prin hol4conc';
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1266
                 message "Modified conclusion:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1267
                 if_debug prin isaconc)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1268
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1269
        fun mk_res th = HOLThm(rens_of info,equal_elim i2h_conc th)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1270
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1271
        case get_hol4_mapping thyname thmname thy of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1272
            SOME (SOME thmname) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1273
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1274
                val th1 = (SOME (PureThy.get_thm thy thmname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1275
                           handle ERROR _ =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1276
                                  (case split_name thmname of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1277
                                       SOME (listname,idx) => (SOME (List.nth(PureThy.get_thms thy listname,idx-1))
28397
389c5e494605 handle _ should be avoided (spurious Interrupt will spoil the game);
wenzelm
parents: 27691
diff changeset
  1278
                                                               handle _ => NONE)  (* FIXME avoid handle _ *)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1279
                                     | NONE => NONE))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1280
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1281
                case th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1282
                    SOME th2 =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1283
                    (case Shuffler.set_prop thy isaconc [(thmname,th2)] of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1284
                         SOME (_,th) => (message "YES";(thy, SOME (mk_res th)))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1285
                       | NONE => (message "NO2";error "get_isabelle_thm" "Bad mapping"))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1286
                  | NONE => (message "NO1";error "get_isabelle_thm" "Bad mapping")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1287
            end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1288
          | SOME NONE => error ("Trying to access ignored theorem " ^ thmname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1289
          | NONE =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1290
            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1291
                val _ = (message "Looking for conclusion:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1292
                         if_debug prin isaconc)
29289
f45c9c3b40a3 eliminated OldTerm.(add_)term_consts;
wenzelm
parents: 29281
diff changeset
  1293
                val cs = non_trivial_term_consts isaconc;
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1294
                val _ = (message "Looking for consts:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1295
                         message (commas cs))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1296
                val pot_thms = Shuffler.find_potential thy isaconc
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1297
                val _ = message ((Int.toString (length pot_thms)) ^ " potential theorems")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1298
            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1299
                case Shuffler.set_prop thy isaconc pot_thms of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1300
                    SOME (isaname,th) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1301
                    let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1302
                        val hth as HOLThm args = mk_res th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1303
                        val thy' =  thy |> add_hol4_theorem thyname thmname args
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1304
                                        |> add_hol4_mapping thyname thmname isaname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1305
                        val _ = ImportRecorder.add_hol_theorem thyname thmname (snd args)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1306
                        val _ = ImportRecorder.add_hol_mapping thyname thmname isaname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1307
                    in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1308
                        (thy',SOME hth)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1309
                    end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1310
                  | NONE => (thy,NONE)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1311
            end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1312
    end
15647
b1f486a9c56b Updated import configuration.
skalberg
parents: 15574
diff changeset
  1313
    handle e => (message "Exception in get_isabelle_thm"; if_debug print_exn e handle _ => (); (thy,NONE))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1314
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1315
fun get_isabelle_thm_and_warn thyname thmname hol4conc thy =
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1316
  let
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1317
    val (a, b) = get_isabelle_thm thyname thmname hol4conc thy
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1318
    fun warn () =
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1319
        let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1320
            val (info,hol4conc') = disamb_term hol4conc
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1321
            val i2h_conc = symmetric (rewrite_hol4_term (HOLogic.mk_Trueprop hol4conc') thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1322
        in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1323
            case concl_of i2h_conc of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1324
                Const("==",_) $ lhs $ _ =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1325
                (warning ("Failed lookup of theorem '"^thmname^"':");
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1326
                 writeln "Original conclusion:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1327
                 prin hol4conc';
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1328
                 writeln "Modified conclusion:";
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1329
                 prin lhs)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1330
              | _ => ()
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1331
        end
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1332
  in
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1333
      case b of
28397
389c5e494605 handle _ should be avoided (spurious Interrupt will spoil the game);
wenzelm
parents: 27691
diff changeset
  1334
          NONE => (warn () handle _ => (); (a,b)) (* FIXME avoid handle _ *)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1335
        | _ => (a, b)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1336
  end
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1337
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1338
fun get_thm thyname thmname thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1339
    case get_hol4_thm thyname thmname thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1340
        SOME hth => (thy,SOME hth)
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1341
      | NONE => ((case import_proof_concl thyname thmname thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1342
                      SOME f => get_isabelle_thm_and_warn thyname thmname (f thy) thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1343
                    | NONE => (message "No conclusion"; (thy,NONE)))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1344
                 handle e as IO.Io _ => (message "IO exception"; (thy,NONE))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1345
                      | e as PK _ => (message "PK exception"; (thy,NONE)))
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1346
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1347
fun rename_const thyname thy name =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1348
    case get_hol4_const_renaming thyname name thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1349
        SOME cname => cname
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15463
diff changeset
  1350
      | NONE => name
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1351
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1352
fun get_def thyname constname rhs thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1353
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1354
        val constname = rename_const thyname thy constname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1355
        val (thmname,thy') = get_defname thyname constname thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1356
        val _ = message ("Looking for definition " ^ thyname ^ "." ^ thmname)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1357
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1358
        get_isabelle_thm_and_warn thyname thmname (mk_teq (thyname ^ "." ^ constname) rhs thy') thy'
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1359
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1360
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1361
fun get_axiom thyname axname thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1362
    case get_thm thyname axname thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1363
        arg as (_,SOME _) => arg
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1364
      | _ => raise ERR "get_axiom" ("Trying to retrieve axiom (" ^ axname ^ ")")
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1365
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1366
fun intern_store_thm gen_output thyname thmname hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1367
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1368
        val (hth' as HOLThm (args as (_,th))) = norm_hthm thy hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1369
        val rew = rewrite_hol4_term (concl_of th) thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1370
        val th  = equal_elim rew th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1371
        val thy' = add_hol4_pending thyname thmname args thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1372
        val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
17657
2f5f595eb618 moved disambiguate_frees to ProofKernel;
wenzelm
parents: 17652
diff changeset
  1373
        val th = disambiguate_frees th
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1374
        val thy2 = if gen_output
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1375
                   then add_dump ("lemma " ^ (quotename thmname) ^ ": " ^
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1376
                                  (smart_string_of_thm th) ^ "\n  by (import " ^
17644
bd59bfd4bf37 fixed disambiguation problem
obua
parents: 17630
diff changeset
  1377
                                  thyname ^ " " ^ (quotename thmname) ^ ")") thy'
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1378
                   else thy'
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1379
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1380
        (thy2,hth')
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1381
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1382
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1383
val store_thm = intern_store_thm true
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1384
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1385
fun mk_REFL ctm =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1386
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1387
        val cty = Thm.ctyp_of_term ctm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1388
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1389
        Drule.instantiate' [SOME cty] [SOME ctm] reflexivity_thm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1390
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1391
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1392
fun REFL tm thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1393
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1394
        val _ = message "REFL:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1395
        val (info,tm') = disamb_term tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1396
        val ctm = Thm.cterm_of thy tm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1397
        val res = HOLThm(rens_of info,mk_REFL ctm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1398
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1399
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1400
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1401
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1402
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1403
fun ASSUME tm thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1404
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1405
        val _ = message "ASSUME:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1406
        val (info,tm') = disamb_term tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1407
        val ctm = Thm.cterm_of thy (HOLogic.mk_Trueprop tm')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1408
        val th = Thm.trivial ctm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1409
        val res = HOLThm(rens_of info,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1410
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1411
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1412
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1413
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1414
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1415
fun INST_TYPE lambda (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1416
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1417
        val _ = message "INST_TYPE:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1418
        val _ = if_debug pth hth
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  1419
        val tys_before = OldTerm.add_term_tfrees (prop_of th,[])
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1420
        val th1 = Thm.varifyT th
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  1421
        val tys_after = OldTerm.add_term_tvars (prop_of th1,[])
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1422
        val tyinst = map (fn (bef, iS) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1423
                             (case try (Lib.assoc (TFree bef)) lambda of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1424
                                  SOME ty => (ctyp_of thy (TVar iS), ctyp_of thy ty)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1425
                                | NONE => (ctyp_of thy (TVar iS), ctyp_of thy (TFree bef))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1426
                             ))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1427
                         (zip tys_before tys_after)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1428
        val res = Drule.instantiate (tyinst,[]) th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1429
        val hth = HOLThm([],res)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1430
        val res = norm_hthm thy hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1431
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1432
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1433
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1434
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1435
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1436
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1437
fun INST sigma hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1438
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1439
        val _ = message "INST:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1440
        val _ = if_debug (app (fn (x,y) => (prin x; prin y))) sigma
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1441
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1442
        val (sdom,srng) = ListPair.unzip (rev sigma)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1443
        val th = hthm2thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1444
        val th1 = mk_INST (map (cterm_of thy) sdom) (map (cterm_of thy) srng) th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1445
        val res = HOLThm([],th1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1446
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1447
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1448
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1449
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1450
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1451
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1452
fun EQ_IMP_RULE (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1453
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1454
        val _ = message "EQ_IMP_RULE:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1455
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1456
        val res = HOLThm(rens,th RS eqimp_thm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1457
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1458
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1459
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1460
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1461
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1462
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  1463
fun mk_EQ_MP th1 th2 = [beta_eta_thm th1, beta_eta_thm th2] MRS eqmp_thm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1464
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1465
fun EQ_MP hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1466
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1467
        val _ = message "EQ_MP:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1468
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1469
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1470
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1471
        val res = HOLThm(rens_of info,mk_EQ_MP th1 th2)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1472
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1473
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1474
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1475
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1476
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1477
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1478
fun mk_COMB th1 th2 thy =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1479
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1480
        val (f,g) = case concl_of th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1481
                        _ $ (Const("op =",_) $ f $ g) => (f,g)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1482
                      | _ => raise ERR "mk_COMB" "First theorem not an equality"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1483
        val (x,y) = case concl_of th2 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1484
                        _ $ (Const("op =",_) $ x $ y) => (x,y)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1485
                      | _ => raise ERR "mk_COMB" "Second theorem not an equality"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1486
        val fty = type_of f
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1487
        val (fd,fr) = dom_rng fty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1488
        val comb_thm' = Drule.instantiate'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1489
                            [SOME (ctyp_of thy fd),SOME (ctyp_of thy fr)]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1490
                            [SOME (cterm_of thy f),SOME (cterm_of thy g),
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1491
                             SOME (cterm_of thy x),SOME (cterm_of thy y)] comb_thm
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1492
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1493
        [th1,th2] MRS comb_thm'
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1494
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1495
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1496
fun SUBST rews ctxt hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1497
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1498
        val _ = message "SUBST:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1499
        val _ = if_debug (app pth) rews
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1500
        val _ = if_debug prin ctxt
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1501
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1502
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1503
        val (info1,ctxt') = disamb_term_from info ctxt
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1504
        val (info2,rews') = disamb_thms_from info1 rews
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1505
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1506
        val cctxt = cterm_of thy ctxt'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1507
        fun subst th [] = th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1508
          | subst th (rew::rews) = subst (mk_COMB th rew thy) rews
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1509
        val res = HOLThm(rens_of info2,mk_EQ_MP (subst (mk_REFL cctxt) rews') th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1510
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1511
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1512
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1513
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1514
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1515
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1516
fun DISJ_CASES hth hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1517
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1518
        val _ = message "DISJ_CASES:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1519
        val _ = if_debug (app pth) [hth,hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1520
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1521
        val (info1,th1) = disamb_thm_from info hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1522
        val (info2,th2) = disamb_thm_from info1 hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1523
        val th1 = norm_hyps th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1524
        val th2 = norm_hyps th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1525
        val (l,r) = case concl_of th of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1526
                        _ $ (Const("op |",_) $ l $ r) => (l,r)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1527
                      | _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1528
        val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1529
        val th2' = rearrange thy (HOLogic.mk_Trueprop r) th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1530
        val res1 = th RS disj_cases_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1531
        val res2 = uniq_compose ((nprems_of th1')-1) th1' ((nprems_of th)+1) res1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1532
        val res3 = uniq_compose ((nprems_of th2')-1) th2' (nprems_of res2) res2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1533
        val res  = HOLThm(rens_of info2,res3)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1534
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1535
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1536
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1537
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1538
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1539
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1540
fun DISJ1 hth tm thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1541
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1542
        val _ = message "DISJ1:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1543
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1544
        val _ = if_debug prin tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1545
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1546
        val (info',tm') = disamb_term_from info tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1547
        val ct = Thm.cterm_of thy tm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1548
        val disj1_thm' = Drule.instantiate' [] [NONE,SOME ct] disj1_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1549
        val res = HOLThm(rens_of info',th RS disj1_thm')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1550
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1551
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1552
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1553
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1554
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1555
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1556
fun DISJ2 tm hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1557
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1558
        val _ = message "DISJ1:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1559
        val _ = if_debug prin tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1560
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1561
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1562
        val (info',tm') = disamb_term_from info tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1563
        val ct = Thm.cterm_of thy tm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1564
        val disj2_thm' = Drule.instantiate' [] [NONE,SOME ct] disj2_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1565
        val res = HOLThm(rens_of info',th RS disj2_thm')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1566
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1567
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1568
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1569
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1570
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1571
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1572
fun IMP_ANTISYM hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1573
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1574
        val _ = message "IMP_ANTISYM:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1575
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1576
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1577
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1578
        val th = [beta_eta_thm th1,beta_eta_thm th2] MRS imp_antisym_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1579
        val res = HOLThm(rens_of info,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1580
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1581
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1582
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1583
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1584
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1585
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1586
fun SYM (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1587
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1588
        val _ = message "SYM:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1589
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1590
        val th = th RS symmetry_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1591
        val res = HOLThm(rens,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1592
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1593
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1594
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1595
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1596
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1597
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1598
fun MP hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1599
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1600
        val _ = message "MP:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1601
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1602
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1603
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1604
        val th = [beta_eta_thm th1,beta_eta_thm th2] MRS mp_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1605
        val res = HOLThm(rens_of info,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1606
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1607
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1608
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1609
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1610
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1611
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1612
fun CONJ hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1613
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1614
        val _ = message "CONJ:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1615
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1616
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1617
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1618
        val th = [th1,th2] MRS conj_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1619
        val res = HOLThm(rens_of info,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1620
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1621
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1622
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1623
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1624
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1625
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1626
fun CONJUNCT1 (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1627
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1628
        val _ = message "CONJUNCT1:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1629
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1630
        val res = HOLThm(rens,th RS conjunct1_thm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1631
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1632
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1633
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1634
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1635
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1636
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1637
fun CONJUNCT2 (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1638
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1639
        val _ = message "CONJUNCT1:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1640
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1641
        val res = HOLThm(rens,th RS conjunct2_thm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1642
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1643
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1644
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1645
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1646
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1647
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1648
fun EXISTS ex wit hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1649
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1650
        val _ = message "EXISTS:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1651
        val _ = if_debug prin ex
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1652
        val _ = if_debug prin wit
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1653
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1654
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1655
        val (info',[ex',wit']) = disamb_terms_from info [ex,wit]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1656
        val cwit = cterm_of thy wit'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1657
        val cty = ctyp_of_term cwit
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1658
        val a = case ex' of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1659
                    (Const("Ex",_) $ a) => a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1660
                  | _ => raise ERR "EXISTS" "Argument not existential"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1661
        val ca = cterm_of thy a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1662
        val exists_thm' = beta_eta_thm (Drule.instantiate' [SOME cty] [SOME ca,SOME cwit] exists_thm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1663
        val th1 = beta_eta_thm th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1664
        val th2 = implies_elim_all th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1665
        val th3 = th2 COMP exists_thm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1666
        val th  = implies_intr_hyps th3
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1667
        val res = HOLThm(rens_of info',th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1668
        val _   = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1669
        val _   = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1670
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1671
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1672
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1673
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1674
fun CHOOSE v hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1675
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1676
        val _ = message "CHOOSE:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1677
        val _ = if_debug prin v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1678
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1679
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1680
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1681
        val (info',v') = disamb_term_from info v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1682
        fun strip 0 _ th = th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1683
          | strip n (p::ps) th =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1684
            strip (n-1) ps (implies_elim th (assume p))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1685
          | strip _ _ _ = raise ERR "CHOOSE" "strip error"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1686
        val cv = cterm_of thy v'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1687
        val th2 = norm_hyps th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1688
        val cvty = ctyp_of_term cv
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1689
        val c = HOLogic.dest_Trueprop (concl_of th2)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1690
        val cc = cterm_of thy c
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1691
        val a = case concl_of th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1692
                    _ $ (Const("Ex",_) $ a) => a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1693
                  | _ => raise ERR "CHOOSE" "Conclusion not existential"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1694
        val ca = cterm_of (theory_of_thm th1) a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1695
        val choose_thm' = beta_eta_thm (Drule.instantiate' [SOME cvty] [SOME ca,SOME cc] choose_thm)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1696
        val th21 = rearrange thy (HOLogic.mk_Trueprop (a $ v')) th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1697
        val th22 = strip ((nprems_of th21)-1) (cprems_of th21) th21
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1698
        val th23 = beta_eta_thm (forall_intr cv th22)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1699
        val th11 = implies_elim_all (beta_eta_thm th1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1700
        val th' = th23 COMP (th11 COMP choose_thm')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1701
        val th = implies_intr_hyps th'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1702
        val res = HOLThm(rens_of info',th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1703
        val _   = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1704
        val _   = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1705
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1706
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1707
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1708
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1709
fun GEN v hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1710
    let
17594
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  1711
      val _ = message "GEN:"
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1712
        val _ = if_debug prin v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1713
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1714
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1715
        val (info',v') = disamb_term_from info v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1716
        val res = HOLThm(rens_of info',mk_GEN v' th thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1717
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1718
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1719
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1720
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1721
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1722
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1723
fun SPEC tm hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1724
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1725
        val _ = message "SPEC:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1726
        val _ = if_debug prin tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1727
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1728
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1729
        val (info',tm') = disamb_term_from info tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1730
        val ctm = Thm.cterm_of thy tm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1731
        val cty = Thm.ctyp_of_term ctm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1732
        val spec' = Drule.instantiate' [SOME cty] [NONE,SOME ctm] spec_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1733
        val th = th RS spec'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1734
        val res = HOLThm(rens_of info',th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1735
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1736
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1737
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1738
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1739
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1740
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1741
fun COMB hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1742
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1743
        val _ = message "COMB:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1744
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1745
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1746
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1747
        val res = HOLThm(rens_of info,mk_COMB th1 th2 thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1748
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1749
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1750
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1751
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1752
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1753
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1754
fun TRANS hth1 hth2 thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1755
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1756
        val _ = message "TRANS:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1757
        val _ = if_debug pth hth1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1758
        val _ = if_debug pth hth2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1759
        val (info,[th1,th2]) = disamb_thms [hth1,hth2]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1760
        val th = [th1,th2] MRS trans_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1761
        val res = HOLThm(rens_of info,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1762
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1763
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1764
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1765
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1766
    end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1767
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1768
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1769
fun CCONTR tm hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1770
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1771
        val _ = message "SPEC:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1772
        val _ = if_debug prin tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1773
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1774
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1775
        val (info',tm') = disamb_term_from info tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1776
        val th = norm_hyps th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1777
        val ct = cterm_of thy tm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1778
        val th1 = rearrange thy (HOLogic.mk_Trueprop (Const("Not",boolT-->boolT) $ tm')) th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1779
        val ccontr_thm' = Drule.instantiate' [] [SOME ct] ccontr_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1780
        val res1 = uniq_compose ((nprems_of th1) - 1) th1 1 ccontr_thm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1781
        val res = HOLThm(rens_of info',res1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1782
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1783
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1784
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1785
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1786
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1787
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1788
fun mk_ABS v th thy =
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1789
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1790
        val cv = cterm_of thy v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1791
        val th1 = implies_elim_all (beta_eta_thm th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1792
        val (f,g) = case concl_of th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1793
                        _ $ (Const("op =",_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1794
                      | _ => raise ERR "mk_ABS" "Bad conclusion"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1795
        val (fd,fr) = dom_rng (type_of f)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1796
        val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1797
        val th2 = forall_intr cv th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1798
        val th3 = th2 COMP abs_thm'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1799
        val res = implies_intr_hyps th3
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1800
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1801
        res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1802
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1803
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1804
fun ABS v hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1805
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1806
        val _ = message "ABS:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1807
        val _ = if_debug prin v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1808
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1809
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1810
        val (info',v') = disamb_term_from info v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1811
        val res = HOLThm(rens_of info',mk_ABS v' th thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1812
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1813
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1814
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1815
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1816
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1817
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1818
fun GEN_ABS copt vlist hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1819
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1820
        val _ = message "GEN_ABS:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1821
        val _ = case copt of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1822
                    SOME c => if_debug prin c
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1823
                  | NONE => ()
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1824
        val _ = if_debug (app prin) vlist
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1825
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1826
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1827
        val (info',vlist') = disamb_terms_from info vlist
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1828
        val th1 =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1829
            case copt of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1830
                SOME (c as Const(cname,cty)) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1831
                let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1832
                    fun inst_type ty1 ty2 (TVar _) = raise ERR "GEN_ABS" "Type variable found!"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1833
                      | inst_type ty1 ty2 (ty as TFree _) = if ty1 = ty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1834
                                                            then ty2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1835
                                                            else ty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1836
                      | inst_type ty1 ty2 (ty as Type(name,tys)) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1837
                        Type(name,map (inst_type ty1 ty2) tys)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1838
                in
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
  1839
                    List.foldr (fn (v,th) =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1840
                              let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1841
                                  val cdom = fst (dom_rng (fst (dom_rng cty)))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1842
                                  val vty  = type_of v
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1843
                                  val newcty = inst_type cdom vty cty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1844
                                  val cc = cterm_of thy (Const(cname,newcty))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1845
                              in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1846
                                  mk_COMB (mk_REFL cc) (mk_ABS v th thy) thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1847
                              end) th vlist'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1848
                end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1849
              | SOME _ => raise ERR "GEN_ABS" "Bad constant"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1850
              | NONE =>
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
  1851
                List.foldr (fn (v,th) => mk_ABS v th thy) th vlist'
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1852
        val res = HOLThm(rens_of info',th1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1853
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1854
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1855
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1856
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1857
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1858
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1859
fun NOT_INTRO (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1860
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1861
        val _ = message "NOT_INTRO:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1862
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1863
        val th1 = implies_elim_all (beta_eta_thm th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1864
        val a = case concl_of th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1865
                    _ $ (Const("op -->",_) $ a $ Const("False",_)) => a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1866
                  | _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1867
        val ca = cterm_of thy a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1868
        val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1869
        val res = HOLThm(rens,implies_intr_hyps th2)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1870
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1871
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1872
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1873
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1874
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1875
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1876
fun NOT_ELIM (hth as HOLThm(rens,th)) thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1877
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1878
        val _ = message "NOT_INTRO:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1879
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1880
        val th1 = implies_elim_all (beta_eta_thm th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1881
        val a = case concl_of th1 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1882
                    _ $ (Const("Not",_) $ a) => a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1883
                  | _ => raise ERR "NOT_ELIM" "Conclusion of bad form"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1884
        val ca = cterm_of thy a
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1885
        val th2 = equal_elim (Drule.instantiate' [] [SOME ca] not_elim_thm) th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1886
        val res = HOLThm(rens,implies_intr_hyps th2)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1887
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1888
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1889
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1890
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1891
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1892
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1893
fun DISCH tm hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1894
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1895
        val _ = message "DISCH:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1896
        val _ = if_debug prin tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1897
        val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1898
        val (info,th) = disamb_thm hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1899
        val (info',tm') = disamb_term_from info tm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1900
        val prems = prems_of th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1901
        val th1 = beta_eta_thm th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1902
        val th2 = implies_elim_all th1
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1903
        val th3 = implies_intr (cterm_of thy (HOLogic.mk_Trueprop tm')) th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1904
        val th4 = th3 COMP disch_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1905
        val res = HOLThm(rens_of info',implies_intr_hyps th4)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1906
        val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1907
        val _ = if_debug pth res
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1908
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1909
        (thy,res)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1910
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1911
32951
83392f9d303f space_implode;
wenzelm
parents: 32945
diff changeset
  1912
val spaces = space_implode " "
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1913
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1914
fun new_definition thyname constname rhs thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1915
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1916
        val constname = rename_const thyname thy constname
17894
f2fdd22accaa Simplifier.theory_context;
wenzelm
parents: 17657
diff changeset
  1917
        val redeclared = isSome (Sign.const_type thy (Sign.intern_const thy constname));
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1918
        val _ = warning ("Introducing constant " ^ constname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1919
        val (thmname,thy) = get_defname thyname constname thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1920
        val (info,rhs') = disamb_term rhs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1921
        val ctype = type_of rhs'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1922
        val csyn = mk_syn thy constname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1923
        val thy1 = case HOL4DefThy.get thy of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1924
                       Replaying _ => thy
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  1925
                     | _ => (ImportRecorder.add_consts [(constname, ctype, csyn)];
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  1926
                              Sign.add_consts_i [(Binding.name constname,ctype,csyn)] thy)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1927
        val eq = mk_defeq constname rhs' thy1
29585
c23295521af5 binding replaces bstring
haftmann
parents: 29289
diff changeset
  1928
        val (thms, thy2) = PureThy.add_defs false [((Binding.name thmname,eq),[])] thy1
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1929
        val _ = ImportRecorder.add_defs thmname eq
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1930
        val def_thm = hd thms
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1931
        val thm' = def_thm RS meta_eq_to_obj_eq_thm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1932
        val (thy',th) = (thy2, thm')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1933
        val fullcname = Sign.intern_const thy' constname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1934
        val thy'' = add_hol4_const_mapping thyname constname true fullcname thy'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1935
        val _ = ImportRecorder.add_hol_const_mapping thyname constname fullcname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1936
        val (linfo,tm24) = disamb_term (mk_teq constname rhs' thy'')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1937
        val rew = rewrite_hol4_term eq thy''
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1938
        val crhs = cterm_of thy'' (#2 (Logic.dest_equals (prop_of rew)))
28677
4693938e9c2a Thm.def_name;
wenzelm
parents: 28662
diff changeset
  1939
        val thy22 = if Thm.def_name constname = thmname andalso not redeclared andalso csyn = NoSyn
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1940
                    then
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1941
                        let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1942
                            val p1 = quotename constname
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1943
                            val p2 = Syntax.string_of_typ_global thy'' ctype
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1944
                            val p3 = string_of_mixfix csyn
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1945
                            val p4 = smart_string_of_cterm crhs
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1946
                        in
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1947
                          add_dump ("constdefs\n  " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ "\n  " ^ p4) thy''
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1948
                        end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1949
                    else
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1950
                        add_dump ("consts\n  " ^ quotename constname ^ " :: \"" ^
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1951
                          Syntax.string_of_typ_global thy'' ctype ^
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1952
                          "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n  " ^
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  1953
                          quotename thmname ^ ": " ^ smart_string_of_cterm crhs) thy''
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1954
        val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1955
                      SOME (_,res) => HOLThm(rens_of linfo,res)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1956
                    | NONE => raise ERR "new_definition" "Bad conclusion"
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28677
diff changeset
  1957
        val fullname = Sign.full_bname thy22 thmname
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1958
        val thy22' = case opt_get_output_thy thy22 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1959
                         "" => (ImportRecorder.add_hol_mapping thyname thmname fullname;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1960
                                add_hol4_mapping thyname thmname fullname thy22)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1961
                       | output_thy =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1962
                         let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1963
                             val moved_thmname = output_thy ^ "." ^ thyname ^ "." ^ thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1964
                             val _ = ImportRecorder.add_hol_move fullname moved_thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1965
                             val _ = ImportRecorder.add_hol_mapping thyname thmname moved_thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1966
                         in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1967
                             thy22 |> add_hol4_move fullname moved_thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1968
                                   |> add_hol4_mapping thyname thmname moved_thmname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1969
                         end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1970
        val _ = message "new_definition:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1971
        val _ = if_debug pth hth
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1972
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1973
        (thy22',hth)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1974
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1975
    handle e => (message "exception in new_definition"; print_exn e)
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1976
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1977
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1978
    val helper = thm "termspec_help"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1979
in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1980
fun new_specification thyname thmname names hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  1981
    case HOL4DefThy.get thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1982
        Replaying _ => (thy,hth)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  1983
      | _ =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1984
        let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1985
            val _ = message "NEW_SPEC:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1986
            val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1987
            val names = map (rename_const thyname thy) names
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1988
            val _ = warning ("Introducing constants " ^ commas names)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1989
            val (HOLThm(rens,th)) = norm_hthm thy hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1990
            val thy1 = case HOL4DefThy.get thy of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1991
                           Replaying _ => thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1992
                         | _ =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1993
                           let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1994
                               fun dest_eta_abs (Abs(x,xT,body)) = (x,xT,body)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1995
                                 | dest_eta_abs body =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1996
                                   let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1997
                                       val (dT,rT) = dom_rng (type_of body)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1998
                                   in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  1999
                                       ("x",dT,body $ Bound 0)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2000
                                   end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2001
                                   handle TYPE _ => raise ERR "new_specification" "not an abstraction type"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2002
                               fun dest_exists (Const("Ex",_) $ abody) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2003
                                   dest_eta_abs abody
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2004
                                 | dest_exists tm =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2005
                                   raise ERR "new_specification" "Bad existential formula"
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2006
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2007
                               val (consts,_) = Library.foldl (fn ((cs,ex),cname) =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2008
                                                          let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2009
                                                              val (_,cT,p) = dest_exists ex
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2010
                                                          in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2011
                                                              ((cname,cT,mk_syn thy cname)::cs,p)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2012
                                                          end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  2013
                               val str = Library.foldl (fn (acc, (c, T, csyn)) =>
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  2014
                                   acc ^ "\n  " ^ quotename c ^ " :: \"" ^
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  2015
                                   Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2016
                               val thy' = add_dump str thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2017
                               val _ = ImportRecorder.add_consts consts
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2018
                           in
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  2019
                               Sign.add_consts_i (map (fn (c, T, mx) => (Binding.name c, T, mx)) consts) thy'
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2020
                           end
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2021
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
  2022
            val thy1 = List.foldr (fn(name,thy)=>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2023
                                snd (get_defname thyname name thy)) thy1 names
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2024
            fun new_name name = fst (get_defname thyname name thy1)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2025
            val names' = map (fn name => (new_name name,name,false)) names
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30447
diff changeset
  2026
            val (thy',res) = Choice_Specification.add_specification NONE
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2027
                                 names'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2028
                                 (thy1,th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2029
            val _ = ImportRecorder.add_specification names' th
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2030
            val res' = Thm.unvarify res
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2031
            val hth = HOLThm(rens,res')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2032
            val rew = rewrite_hol4_term (concl_of res') thy'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2033
            val th  = equal_elim rew res'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2034
            fun handle_const (name,thy) =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2035
                let
28677
4693938e9c2a Thm.def_name;
wenzelm
parents: 28662
diff changeset
  2036
                    val defname = Thm.def_name name
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2037
                    val (newname,thy') = get_defname thyname name thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2038
                in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2039
                    (if defname = newname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2040
                     then quotename name
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2041
                     else (quotename newname) ^ ": " ^ (quotename name),thy')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2042
                end
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 29585
diff changeset
  2043
            val (new_names,thy') = List.foldr (fn(name,(names,thy)) =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2044
                                            let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2045
                                                val (name',thy') = handle_const (name,thy)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2046
                                            in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2047
                                                (name'::names,thy')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2048
                                            end) ([],thy') names
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2049
            val thy'' = add_dump ("specification (" ^ (spaces new_names) ^ ") " ^ thmname ^ ": " ^ (smart_string_of_thm th) ^
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2050
                                  "\n  by (import " ^ thyname ^ " " ^ thmname ^ ")")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2051
                                 thy'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2052
            val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2053
            val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2054
        in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2055
            intern_store_thm false thyname thmname hth thy''
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2056
        end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2057
        handle e => (message "exception in new_specification"; print_exn e)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2058
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2059
end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2060
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2061
fun new_axiom name tm thy = raise ERR "new_axiom" ("Oh, no you don't! (" ^ name ^ ")")
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2062
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2063
fun to_isa_thm (hth as HOLThm(_,th)) =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2064
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2065
        val (HOLThm args) = norm_hthm (theory_of_thm th) hth
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2066
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2067
        apsnd strip_shyps args
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2068
    end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2069
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2070
fun to_isa_term tm = tm
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2071
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2072
local
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2073
    val light_nonempty = thm "light_ex_imp_nonempty"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2074
    val ex_imp_nonempty = thm "ex_imp_nonempty"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2075
    val typedef_hol2hol4 = thm "typedef_hol2hol4"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2076
    val typedef_hol2hollight = thm "typedef_hol2hollight"
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2077
in
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2078
fun new_type_definition thyname thmname tycname hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2079
    case HOL4DefThy.get thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2080
        Replaying _ => (thy,hth)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2081
      | _ =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2082
        let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2083
            val _ = message "TYPE_DEF:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2084
            val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2085
            val _ = warning ("Introducing type " ^ tycname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2086
            val (HOLThm(rens,td_th)) = norm_hthm thy hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2087
            val th2 = beta_eta_thm (td_th RS ex_imp_nonempty)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2088
            val c = case concl_of th2 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2089
                        _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2090
                      | _ => raise ERR "new_type_definition" "Bad type definition theorem"
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  2091
            val tfrees = OldTerm.term_tfrees c
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2092
            val tnames = map fst tfrees
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2093
            val tsyn = mk_syn thy tycname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2094
            val typ = (tycname,tnames,tsyn)
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  2095
            val ((_, typedef_info), thy') =
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30447
diff changeset
  2096
              Typedef.add_typedef false (SOME (Binding.name thmname))
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  2097
                (Binding.name tycname, tnames, tsyn) c NONE (rtac th2 1) thy
19064
bf19cc5a7899 fixed bugs, added caching
obua
parents: 18929
diff changeset
  2098
            val _ = ImportRecorder.add_typedef (SOME thmname) typ c NONE th2
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2099
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2100
            val th3 = (#type_definition typedef_info) RS typedef_hol2hol4
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2101
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2102
            val fulltyname = Sign.intern_type thy' tycname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2103
            val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2104
            val _ = ImportRecorder.add_hol_type_mapping thyname tycname fulltyname
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2105
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2106
            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th3))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2107
            val _ = if has_ren hth' then warning ("Theorem " ^ thmname ^ " needs variable-disambiguating")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2108
                    else ()
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2109
            val thy4 = add_hol4_pending thyname thmname args thy''
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2110
            val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2111
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2112
            val rew = rewrite_hol4_term (concl_of td_th) thy4
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2113
            val th  = equal_elim rew (Thm.transfer thy4 td_th)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2114
            val c   = case HOLogic.dest_Trueprop (prop_of th) of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2115
                          Const("Ex",exT) $ P =>
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2116
                          let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2117
                              val PT = domain_type exT
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2118
                          in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2119
                              Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2120
                          end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2121
                        | _ => error "Internal error in ProofKernel.new_typedefinition"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2122
            val tnames_string = if null tnames
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2123
                                then ""
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2124
                                else "(" ^ commas tnames ^ ") "
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2125
            val proc_prop = if null tnames
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2126
                            then smart_string_of_cterm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2127
                            else Library.setmp show_all_types true smart_string_of_cterm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2128
            val thy5 = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^ " = " ^ (proc_prop (cterm_of thy4 c)) ^ " "
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2129
                                 ^ (string_of_mixfix tsyn) ^ "\n  by (rule typedef_helper,import " ^ thyname ^ " " ^ thmname ^ ")") thy4
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2130
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2131
            val thy6 = add_dump ("lemmas " ^ thmname ^ " = typedef_hol2hol4 [OF type_definition_" ^ tycname ^ "]") thy5
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2132
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2133
            val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2134
            val _ = if_debug pth hth'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2135
        in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2136
            (thy6,hth')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2137
        end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2138
        handle e => (message "exception in new_type_definition"; print_exn e)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2139
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2140
fun add_dump_constdefs thy defname constname rhs ty =
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2141
    let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2142
        val n = quotename constname
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  2143
        val t = Syntax.string_of_typ_global thy ty
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2144
        val syn = string_of_mixfix (mk_syn thy constname)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2145
        (*val eq = smart_string_of_cterm (cterm_of thy (Const(rhs, ty)))*)
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2146
        val eq = quote (constname ^ " == "^rhs)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2147
        val d = case defname of NONE => "" | SOME defname => (quotename defname)^" : "
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2148
    in
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2149
        add_dump ("constdefs\n  " ^ n ^ " :: \"" ^ t ^ "\" " ^ syn ^ "\n  " ^ d ^ eq) thy
17490
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2150
    end
ec62f340e811 maybe the last bug fix (sigh)?
obua
parents: 17444
diff changeset
  2151
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2152
fun add_dump_syntax thy name =
17594
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2153
    let
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2154
      val n = quotename name
19264
61e775c03ed8 string_of_mixfix;
wenzelm
parents: 19068
diff changeset
  2155
      val syn = string_of_mixfix (mk_syn thy name)
17594
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2156
    in
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2157
      add_dump ("syntax\n  "^n^" :: _ "^syn) thy
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2158
    end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2159
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2160
fun type_introduction thyname thmname tycname abs_name rep_name (P,t) hth thy =
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2161
    case HOL4DefThy.get thy of
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2162
        Replaying _ => (thy,
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
  2163
          HOLThm([], PureThy.get_thm thy (thmname^"_@intern")) handle ERROR _ => hth)
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2164
      | _ =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2165
        let
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  2166
            val _ = message "TYPE_INTRO:"
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2167
            val _ = if_debug pth hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2168
            val _ = warning ("Introducing type " ^ tycname ^ " (with morphisms " ^ abs_name ^ " and " ^ rep_name ^ ")")
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2169
            val (HOLThm(rens,td_th)) = norm_hthm thy hth
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2170
            val tT = type_of t
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2171
            val light_nonempty' =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2172
                Drule.instantiate' [SOME (ctyp_of thy tT)]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2173
                                   [SOME (cterm_of thy P),
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2174
                                    SOME (cterm_of thy t)] light_nonempty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2175
            val th2 = beta_eta_thm (td_th RS (beta_eta_thm light_nonempty'))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2176
            val c = case concl_of th2 of
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2177
                        _ $ (Const("Ex",_) $ Abs(_,_,Const("op :",_) $ _ $ c)) => c
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2178
                      | _ => raise ERR "type_introduction" "Bad type definition theorem"
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
  2179
            val tfrees = OldTerm.term_tfrees c
32945
63db9da65a19 sort_strings (cf. Pure/library.ML);
wenzelm
parents: 32740
diff changeset
  2180
            val tnames = sort_strings (map fst tfrees)
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2181
            val tsyn = mk_syn thy tycname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2182
            val typ = (tycname,tnames,tsyn)
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  2183
            val ((_, typedef_info), thy') =
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30447
diff changeset
  2184
              Typedef.add_typedef false NONE (Binding.name tycname,tnames,tsyn) c
30346
90efbb8a8cb2 minimal adaptions for abstract binding type;
wenzelm
parents: 30190
diff changeset
  2185
                (SOME(Binding.name rep_name,Binding.name abs_name)) (rtac th2 1) thy
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2186
            val _ = ImportRecorder.add_typedef NONE typ c (SOME(rep_name,abs_name)) th2
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2187
            val fulltyname = Sign.intern_type thy' tycname
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2188
            val aty = Type (fulltyname, map mk_vartype tnames)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2189
            val abs_ty = tT --> aty
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2190
            val rep_ty = aty --> tT
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2191
            val typedef_hol2hollight' =
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2192
                Drule.instantiate'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2193
                    [SOME (ctyp_of thy' aty), SOME (ctyp_of thy' tT)]
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2194
                    [NONE, NONE, NONE, SOME (cterm_of thy' (Free ("a", aty))), SOME (cterm_of thy' (Free ("r", tT)))]
17379
85109eec887b Fixed Importer bug in type_introduction: instantiate type variables in rep-abs theorems exactly as it is done in HOL-light.
obua
parents: 17335
diff changeset
  2195
                    typedef_hol2hollight
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2196
            val th4 = (#type_definition typedef_info) RS typedef_hol2hollight'
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22675
diff changeset
  2197
            val _ = null (Thm.fold_terms Term.add_tvars th4 []) orelse
20286
4cf8e86a2d29 removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 19998
diff changeset
  2198
              raise ERR "type_introduction" "no type variables expected any more"
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22675
diff changeset
  2199
            val _ = null (Thm.fold_terms Term.add_vars th4 []) orelse
20286
4cf8e86a2d29 removed obsolete Drule.frees/vars_of etc.;
wenzelm
parents: 19998
diff changeset
  2200
              raise ERR "type_introduction" "no term variables expected any more"
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2201
            val _ = message ("step 3: thyname="^thyname^", tycname="^tycname^", fulltyname="^fulltyname)
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  2202
            val thy'' = add_hol4_type_mapping thyname tycname true fulltyname thy'
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2203
            val _ = ImportRecorder.add_hol_type_mapping thyname tycname fulltyname
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  2204
            val _ = message "step 4"
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2205
            val (hth' as HOLThm args) = norm_hthm thy'' (HOLThm(rens,th4))
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2206
            val thy4 = add_hol4_pending thyname thmname args thy''
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2207
            val _ = ImportRecorder.add_hol_pending thyname thmname (hthm2thm hth')
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2208
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2209
            val P' = P (* why !? #2 (Logic.dest_equals (concl_of (rewrite_hol4_term P thy4))) *)
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2210
            val c   =
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2211
                let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2212
                    val PT = type_of P'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2213
                in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2214
                    Const("Collect",PT-->HOLogic.mk_setT (domain_type PT)) $ P'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2215
                end
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2216
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2217
            val tnames_string = if null tnames
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2218
                                then ""
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2219
                                else "(" ^ commas tnames ^ ") "
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2220
            val proc_prop = if null tnames
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2221
                            then smart_string_of_cterm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2222
                            else Library.setmp show_all_types true smart_string_of_cterm
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2223
            val thy = add_dump ("typedef (open) " ^ tnames_string ^ (quotename tycname) ^
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2224
              " = " ^ (proc_prop (cterm_of thy4 c)) ^ " " ^
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2225
              (string_of_mixfix tsyn) ^ " morphisms "^
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2226
              (quote rep_name)^" "^(quote abs_name)^"\n"^
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2227
              ("  apply (rule light_ex_imp_nonempty[where t="^
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2228
              (proc_prop (cterm_of thy4 t))^"])\n"^
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2229
              ("  by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32180
diff changeset
  2230
            val str_aty = Syntax.string_of_typ_global thy aty
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24634
diff changeset
  2231
            val thy = add_dump_syntax thy rep_name
17594
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2232
            val thy = add_dump_syntax thy abs_name
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2233
            val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
17594
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2234
              " = typedef_hol2hollight \n"^
98be710dabc4 replay type_introduction fix
obua
parents: 17592
diff changeset
  2235
              "  [where a=\"a :: "^str_aty^"\" and r=r" ^
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2236
              " ,\n   OF "^(quotename ("type_definition_" ^ tycname)) ^ "]") thy
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2237
            val _ = message "RESULT:"
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2238
            val _ = if_debug pth hth'
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2239
        in
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2240
            (thy,hth')
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2241
        end
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26343
diff changeset
  2242
        handle e => (message "exception in type_introduction"; print_exn e)
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2243
end
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2244
17322
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  2245
val prin = prin
781abf7011e6 Added HOLLight support to importer.
obua
parents: 16486
diff changeset
  2246
14516
a183dec876ab Added HOL proof importer.
skalberg
parents:
diff changeset
  2247
end