src/Pure/zterm.ML
author wenzelm
Wed, 06 Dec 2023 16:04:00 +0100
changeset 79152 4189e10f1524
parent 79150 1cdc685fe852
child 79153 16a144eaf67d
permissions -rw-r--r--
more zproofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/zterm.ML
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     3
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     4
Tight representation of types / terms / proof terms, notably for proof recording.
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     5
*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
     6
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
     7
(*** global ***)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
     8
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
     9
(* types and terms *)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    10
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    11
datatype ztyp =
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    12
    ZTVar of indexname * sort      (*free: index ~1*)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    13
  | ZFun of ztyp * ztyp
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    14
  | ZProp
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    15
  | ZItself of ztyp
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    16
  | ZType0 of string               (*type constant*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    17
  | ZType1 of string * ztyp        (*type constructor: 1 argument*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    18
  | ZType of string * ztyp list    (*type constructor: >= 2 arguments*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    19
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    20
datatype zterm =
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    21
    ZVar of indexname * ztyp       (*free: index ~1*)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    22
  | ZBound of int
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    23
  | ZConst0 of string              (*monomorphic constant*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    24
  | ZConst1 of string * ztyp       (*polymorphic constant: 1 type argument*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    25
  | ZConst of string * ztyp list   (*polymorphic constant: >= 2 type arguments*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    26
  | ZAbs of string * ztyp * zterm
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    27
  | ZApp of zterm * zterm
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    28
  | ZClass of ztyp * class         (*OFCLASS proposition*)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    29
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    30
structure ZTerm =
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    31
struct
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    32
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    33
(* fold *)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    34
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    35
fun fold_tvars f (ZTVar v) = f v
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    36
  | fold_tvars f (ZFun (T, U)) = fold_tvars f T #> fold_tvars f U
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    37
  | fold_tvars f (ZItself T) = fold_tvars f T
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    38
  | fold_tvars f (ZType1 (_, T)) = fold_tvars f T
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    39
  | fold_tvars f (ZType (_, Ts)) = fold (fold_tvars f) Ts
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    40
  | fold_tvars _ _ = I;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    41
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    42
fun fold_aterms f (ZApp (t, u)) = fold_aterms f t #> fold_aterms f u
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    43
  | fold_aterms f (ZAbs (_, _, t)) = fold_aterms f t
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    44
  | fold_aterms f a = f a;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    45
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    46
fun fold_types f (ZVar (_, T)) = f T
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    47
  | fold_types f (ZConst1 (_, T)) = f T
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    48
  | fold_types f (ZConst (_, As)) = fold f As
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    49
  | fold_types f (ZAbs (_, T, b)) = f T #> fold_types f b
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    50
  | fold_types f (ZApp (t, u)) = fold_types f t #> fold_types f u
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    51
  | fold_types f (ZClass (T, _)) = f T
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    52
  | fold_types _ _ = I;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    53
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    54
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    55
(* ordering *)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    56
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    57
local
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    58
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    59
fun cons_nr (ZTVar _) = 0
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    60
  | cons_nr (ZFun _) = 1
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    61
  | cons_nr ZProp = 2
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    62
  | cons_nr (ZItself _) = 3
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    63
  | cons_nr (ZType0 _) = 4
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    64
  | cons_nr (ZType1 _) = 5
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    65
  | cons_nr (ZType _) = 6;
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    66
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    67
val fast_indexname_ord = Term_Ord.fast_indexname_ord;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    68
val sort_ord = Term_Ord.sort_ord;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    69
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    70
in
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    71
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    72
fun ztyp_ord TU =
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    73
  if pointer_eq TU then EQUAL
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    74
  else
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    75
    (case TU of
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    76
      (ZTVar (a, A), ZTVar (b, B)) =>
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
    77
        (case fast_indexname_ord (a, b) of EQUAL => sort_ord (A, B) | ord => ord)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    78
    | (ZFun (T, T'), ZFun (U, U')) =>
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    79
        (case ztyp_ord (T, U) of EQUAL => ztyp_ord (T', U') | ord => ord)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    80
    | (ZProp, ZProp) => EQUAL
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    81
    | (ZItself T, ZItself U) => ztyp_ord (T, U)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    82
    | (ZType0 a, ZType0 b) => fast_string_ord (a, b)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    83
    | (ZType1 (a, T), ZType1 (b, U)) =>
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    84
        (case fast_string_ord (a, b) of EQUAL => ztyp_ord (T, U) | ord => ord)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    85
    | (ZType (a, Ts), ZType (b, Us)) =>
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    86
        (case fast_string_ord (a, b) of EQUAL => dict_ord ztyp_ord (Ts, Us) | ord => ord)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    87
    | (T, U) => int_ord (cons_nr T, cons_nr U));
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    88
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    89
end;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    90
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    91
end;
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
    92
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    93
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    94
(* term items *)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    95
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    96
structure ZTVars:
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    97
sig
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    98
  include TERM_ITEMS
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
    99
  val add_tvarsT: ztyp -> set -> set
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   100
  val add_tvars: zterm -> set -> set
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   101
end =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   102
struct
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   103
  open TVars;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   104
  val add_tvarsT = ZTerm.fold_tvars add_set;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   105
  val add_tvars = ZTerm.fold_types add_tvarsT;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   106
end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   107
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   108
structure ZVars:
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   109
sig
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   110
  include TERM_ITEMS
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   111
  val add_vars: zterm -> set -> set
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   112
end =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   113
struct
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   114
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   115
structure Term_Items = Term_Items
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   116
(
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   117
  type key = indexname * ztyp;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   118
  val ord = pointer_eq_ord (prod_ord Term_Ord.fast_indexname_ord ZTerm.ztyp_ord);
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   119
);
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   120
open Term_Items;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   121
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   122
val add_vars = ZTerm.fold_aterms (fn ZVar v => add_set v | _ => I);
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   123
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   124
end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   125
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   126
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   127
(* proofs *)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   128
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   129
datatype zproof_name =
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   130
    ZAxiom of string
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   131
  | ZOracle of string
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   132
  | ZBox of serial;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   133
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   134
datatype zproof =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   135
    ZDummy                         (*dummy proof*)
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   136
  | ZConstP of zproof_name * zterm * ztyp ZTVars.table * zterm ZVars.table
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   137
  | ZBoundP of int
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   138
  | ZHyp of zterm
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   139
  | ZAbst of string * ztyp * zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   140
  | ZAbsP of string * zterm * zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   141
  | ZAppt of zproof * zterm
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   142
  | ZAppP of zproof * zproof
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   143
  | ZClassP of ztyp * class;       (*OFCLASS proof from sorts algebra*)
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   144
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   145
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   146
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   147
(*** local ***)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   148
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   149
signature ZTERM =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   150
sig
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   151
  datatype ztyp = datatype ztyp
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   152
  datatype zterm = datatype zterm
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   153
  datatype zproof = datatype zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   154
  val fold_tvars: (indexname * sort -> 'a -> 'a) -> ztyp -> 'a -> 'a
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   155
  val fold_aterms: (zterm -> 'a -> 'a) -> zterm -> 'a -> 'a
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   156
  val fold_types: (ztyp -> 'a -> 'a) -> zterm -> 'a -> 'a
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   157
  val ztyp_ord: ztyp * ztyp -> order
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   158
  val aconv_zterm: zterm * zterm -> bool
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   159
  val ztyp_of: typ -> ztyp
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   160
  val typ_of: ztyp -> typ
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   161
  val zterm_of: Consts.T -> term -> zterm
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   162
  val term_of: Consts.T -> zterm -> term
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   163
  val global_zterm_of: theory -> term -> zterm
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   164
  val global_term_of: theory -> zterm -> term
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   165
  val dummy_proof: 'a -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   166
  val todo_proof: 'a -> zproof
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   167
  val axiom_proof:  theory -> string -> term -> zproof
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   168
  val oracle_proof:  theory -> string -> term -> zproof
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   169
  val assume_proof: theory -> term -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   170
  val trivial_proof: theory -> term -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   171
  val implies_intr_proof: theory -> term -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   172
  val forall_intr_proof: theory -> typ -> string * term -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   173
  val forall_elim_proof: theory -> term -> zproof -> zproof
79128
b6f5d4392388 more zproofs;
wenzelm
parents: 79126
diff changeset
   174
  val of_class_proof: typ * class -> zproof
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   175
  val reflexive_proof: theory -> typ -> term -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   176
  val symmetric_proof: theory -> typ -> term -> term -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   177
  val transitive_proof: theory -> typ -> term -> term -> term -> zproof -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   178
  val equal_intr_proof: theory -> term -> term -> zproof -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   179
  val equal_elim_proof: theory -> term -> term -> zproof -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   180
  val abstract_rule_proof: theory -> typ -> typ -> string * term -> term -> term -> zproof -> zproof
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   181
  val combination_proof: theory -> typ -> typ -> term -> term -> term -> term ->
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   182
    zproof -> zproof -> zproof
79133
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   183
  val generalize_proof: Names.set * Names.set -> int -> zproof -> zproof
79149
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   184
  val instantiate_proof: theory ->
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   185
    ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> zproof -> zproof
79135
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   186
  val varifyT_proof: ((string * sort) * (indexname * sort)) list -> zproof -> zproof
79152
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   187
  val legacy_freezeT_proof: term -> zproof -> zproof
79150
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   188
  val incr_indexes_proof: int -> zproof -> zproof
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   189
end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   190
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   191
structure ZTerm: ZTERM =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   192
struct
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   193
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   194
datatype ztyp = datatype ztyp;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   195
datatype zterm = datatype zterm;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   196
datatype zproof = datatype zproof;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   197
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   198
open ZTerm;
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   199
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   200
fun aconv_zterm (tm1, tm2) =
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   201
  pointer_eq (tm1, tm2) orelse
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   202
    (case (tm1, tm2) of
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   203
      (ZApp (t1, u1), ZApp (t2, u2)) => aconv_zterm (t1, t2) andalso aconv_zterm (u1, u2)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   204
    | (ZAbs (_, T1, t1), ZAbs (_, T2, t2)) => aconv_zterm (t1, t2) andalso T1 = T2
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   205
    | (a1, a2) => a1 = a2);
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   206
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   207
79132
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   208
(* map structure *)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   209
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   210
fun subst_type_same tvar =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   211
  let
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   212
    fun typ (ZTVar x) = tvar x
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   213
      | typ (ZFun (T, U)) = (ZFun (typ T, Same.commit typ U) handle Same.SAME => ZFun (T, typ U))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   214
      | typ ZProp = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   215
      | typ (ZItself T) = ZItself (typ T)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   216
      | typ (ZType0 _) = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   217
      | typ (ZType1 (a, T)) = ZType1 (a, typ T)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   218
      | typ (ZType (a, Ts)) = ZType (a, Same.map typ Ts);
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   219
  in typ end;
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   220
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   221
fun subst_term_same typ var =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   222
  let
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   223
    fun term (ZVar (x, T)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   224
          let val (T', same) = Same.commit_id typ T in
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   225
            (case Same.catch var (x, T') of
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   226
              NONE => if same then raise Same.SAME else ZVar (x, T')
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   227
            | SOME t' => t')
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   228
          end
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   229
      | term (ZBound _) = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   230
      | term (ZConst0 _) = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   231
      | term (ZConst1 (a, T)) = ZConst1 (a, typ T)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   232
      | term (ZConst (a, Ts)) = ZConst (a, Same.map typ Ts)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   233
      | term (ZAbs (a, T, t)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   234
          (ZAbs (a, typ T, Same.commit term t) handle Same.SAME => ZAbs (a, T, term t))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   235
      | term (ZApp (t, u)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   236
          (ZApp (term t, Same.commit term u) handle Same.SAME => ZApp (t, term u))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   237
      | term (ZClass (T, c)) = ZClass (typ T, c);
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   238
  in term end;
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   239
79145
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   240
fun map_insts_same typ term (instT, inst) =
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   241
  let
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   242
    val changed = Unsynchronized.ref false;
79146
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   243
    fun apply f x =
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   244
      (case Same.catch f x of
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   245
        NONE => NONE
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   246
      | some => (changed := true; some));
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   247
79145
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   248
    val instT' =
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   249
      (instT, instT) |-> ZTVars.fold (fn (v, T) =>
79146
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   250
        (case apply typ T of
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   251
          NONE => I
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   252
        | SOME T' => ZTVars.update (v, T')));
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   253
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   254
    val vars' =
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   255
      (inst, ZVars.empty) |-> ZVars.fold (fn ((v, T), _) =>
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   256
        (case apply typ T of
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   257
          NONE => I
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   258
        | SOME T' => ZVars.add ((v, T), (v, T'))));
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   259
79145
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   260
    val inst' =
79146
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   261
      if ZVars.is_empty vars' then
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   262
        (inst, inst) |-> ZVars.fold (fn (v, t) =>
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   263
          (case apply term t of
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   264
            NONE => I
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   265
          | SOME t' => ZVars.update (v, t')))
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   266
      else
79145
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   267
        ZVars.dest inst
79146
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   268
        |> map (fn (v, t) => (the_default v (ZVars.lookup vars' v), the_default t (apply term t)))
feb94ac5df41 more accurate treatment of term variables after instantiation of type variables;
wenzelm
parents: 79145
diff changeset
   269
        |> ZVars.make_strict;
79145
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   270
  in if ! changed then (instT', inst') else raise Same.SAME end;
a9c55fef42b0 tuned signature;
wenzelm
parents: 79144
diff changeset
   271
79132
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   272
fun map_proof_same typ term =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   273
  let
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   274
    fun proof ZDummy = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   275
      | proof (ZConstP (a, A, instT, inst)) =
79148
99201e7b1d94 proper treatment of ZConstP: term represents body of closure;
wenzelm
parents: 79147
diff changeset
   276
          let val (instT', inst') = map_insts_same typ term (instT, inst)
99201e7b1d94 proper treatment of ZConstP: term represents body of closure;
wenzelm
parents: 79147
diff changeset
   277
          in ZConstP (a, A, instT', inst') end
79132
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   278
      | proof (ZBoundP _) = raise Same.SAME
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   279
      | proof (ZHyp h) = ZHyp (term h)
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   280
      | proof (ZAbst (a, T, p)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   281
          (ZAbst (a, typ T, Same.commit proof p) handle Same.SAME => ZAbst (a, T, proof p))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   282
      | proof (ZAbsP (a, t, p)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   283
          (ZAbsP (a, term t, Same.commit proof p) handle Same.SAME => ZAbsP (a, t, proof p))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   284
      | proof (ZAppt (p, t)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   285
          (ZAppt (proof p, Same.commit term t) handle Same.SAME => ZAppt (p, term t))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   286
      | proof (ZAppP (p, q)) =
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   287
          (ZAppP (proof p, Same.commit proof q) handle Same.SAME => ZAppP (p, proof q))
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   288
      | proof (ZClassP (T, c)) = ZClassP (typ T, c);
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   289
  in proof end;
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   290
79144
wenzelm
parents: 79136
diff changeset
   291
fun map_proof_types_same typ =
wenzelm
parents: 79136
diff changeset
   292
  map_proof_same typ (subst_term_same typ Same.same);
wenzelm
parents: 79136
diff changeset
   293
79132
6d3322477cfd more operations;
wenzelm
parents: 79128
diff changeset
   294
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   295
(* instantiation *)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   296
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   297
fun init_instT t = ZTVars.build (ZTVars.add_tvars t) |> ZTVars.map (fn v => fn _ => ZTVar v);
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   298
fun init_inst t = ZVars.build (ZVars.add_vars t) |> ZVars.map (fn v => fn _ => ZVar v);
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   299
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   300
fun map_const_proof (f, g) prf =
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   301
  (case prf of
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   302
    ZConstP (a, A, instT, inst) =>
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   303
      let
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   304
        val instT' = ZTVars.map (fn ((x, _), _) => fn y => the_default y (try f x)) instT;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   305
        val inst' = ZVars.map (fn ((x, _), _) => fn y => the_default y (try g x)) inst;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   306
      in ZConstP (a, A, instT', inst') end
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   307
  | _ => prf);
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   308
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   309
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   310
(* convert ztyp / zterm vs. regular typ / term *)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   311
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   312
fun ztyp_of (TFree (a, S)) = ZTVar ((a, ~1), S)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   313
  | ztyp_of (TVar v) = ZTVar v
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   314
  | ztyp_of (Type ("fun", [T, U])) = ZFun (ztyp_of T, ztyp_of U)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   315
  | ztyp_of (Type (c, [])) = if c = "prop" then ZProp else ZType0 c
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   316
  | ztyp_of (Type (c, [T])) = if c = "itself" then ZItself (ztyp_of T) else ZType1 (c, ztyp_of T)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   317
  | ztyp_of (Type (c, ts)) = ZType (c, map ztyp_of ts);
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   318
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   319
fun typ_of (ZTVar ((a, ~1), S)) = TFree (a, S)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   320
  | typ_of (ZTVar v) = TVar v
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   321
  | typ_of (ZFun (T, U)) = typ_of T --> typ_of U
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   322
  | typ_of ZProp = propT
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   323
  | typ_of (ZItself T) = Term.itselfT (typ_of T)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   324
  | typ_of (ZType0 c) = Type (c, [])
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   325
  | typ_of (ZType1 (c, T)) = Type (c, [typ_of T])
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   326
  | typ_of (ZType (c, Ts)) = Type (c, map typ_of Ts);
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   327
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   328
fun zterm_of consts =
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   329
  let
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   330
    val typargs = Consts.typargs consts;
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   331
    fun zterm (Free (x, T)) = ZVar ((x, ~1), ztyp_of T)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   332
      | zterm (Var (xi, T)) = ZVar (xi, ztyp_of T)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   333
      | zterm (Bound i) = ZBound i
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   334
      | zterm (Const (c, T)) =
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   335
          (case typargs (c, T) of
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   336
            [] => ZConst0 c
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   337
          | [T] => ZConst1 (c, ztyp_of T)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   338
          | Ts => ZConst (c, map ztyp_of Ts))
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   339
      | zterm (Abs (a, T, b)) = ZAbs (a, ztyp_of T, zterm b)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   340
      | zterm ((t as Const (c, _)) $ (u as Const ("Pure.type", _))) =
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   341
          if String.isSuffix Logic.class_suffix c then
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   342
            ZClass (ztyp_of (Logic.dest_type u), Logic.class_of_const c)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   343
          else ZApp (zterm t, zterm u)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   344
      | zterm (t $ u) = ZApp (zterm t, zterm u);
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   345
  in zterm end;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   346
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   347
fun term_of consts =
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   348
  let
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   349
    val instance = Consts.instance consts;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   350
    fun const (c, Ts) = Const (c, instance (c, Ts));
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   351
    fun term (ZVar ((x, ~1), T)) = Free (x, typ_of T)
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   352
      | term (ZVar (xi, T)) = Var (xi, typ_of T)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   353
      | term (ZBound i) = Bound i
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   354
      | term (ZConst0 c) = const (c, [])
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   355
      | term (ZConst1 (c, T)) = const (c, [typ_of T])
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   356
      | term (ZConst (c, Ts)) = const (c, map typ_of Ts)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   357
      | term (ZAbs (a, T, b)) = Abs (a, typ_of T, term b)
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   358
      | term (ZApp (t, u)) = term t $ term u
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   359
      | term (ZClass (T, c)) = Logic.mk_of_class (typ_of T, c);
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   360
  in term end;
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   361
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   362
val global_zterm_of = zterm_of o Sign.consts_of;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   363
val global_term_of = term_of o Sign.consts_of;
79113
5109e4b2a292 pro-forma support for optional zproof: no proper content yet;
wenzelm
parents: 79098
diff changeset
   364
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   365
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   366
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   367
(** proof construction **)
79113
5109e4b2a292 pro-forma support for optional zproof: no proper content yet;
wenzelm
parents: 79098
diff changeset
   368
5109e4b2a292 pro-forma support for optional zproof: no proper content yet;
wenzelm
parents: 79098
diff changeset
   369
fun dummy_proof _ = ZDummy;
5109e4b2a292 pro-forma support for optional zproof: no proper content yet;
wenzelm
parents: 79098
diff changeset
   370
val todo_proof = dummy_proof;
5109e4b2a292 pro-forma support for optional zproof: no proper content yet;
wenzelm
parents: 79098
diff changeset
   371
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   372
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   373
(* basic logic *)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   374
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   375
fun const_proof thy a A =
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   376
  let
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   377
    val t = global_zterm_of thy A;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   378
    val instT = init_instT t;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   379
    val inst = init_inst t;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   380
  in ZConstP (a, t, instT, inst) end;
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   381
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   382
fun axiom_proof thy name = const_proof thy (ZAxiom name);
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   383
fun oracle_proof thy name = const_proof thy (ZOracle name);
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   384
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   385
fun assume_proof thy A =
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   386
  ZHyp (global_zterm_of thy A);
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   387
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   388
fun trivial_proof thy A =
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   389
  ZAbsP ("H", global_zterm_of thy A, ZBoundP 0);
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   390
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   391
fun implies_intr_proof thy A prf =
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   392
  let
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   393
    val h = global_zterm_of thy A;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   394
    fun abs_hyp i (p as ZHyp t) = if aconv_zterm (h, t) then ZBoundP i else p
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   395
      | abs_hyp i (ZAbst (x, T, p)) = ZAbst (x, T, abs_hyp i p)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   396
      | abs_hyp i (ZAbsP (x, t, p)) = ZAbsP (x, t, abs_hyp (i + 1) p)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   397
      | abs_hyp i (ZAppt (p, t)) = ZAppt (abs_hyp i p, t)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   398
      | abs_hyp i (ZAppP (p, q)) = ZAppP (abs_hyp i p, abs_hyp i q)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   399
      | abs_hyp _ p = p;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   400
  in ZAbsP ("H", h, abs_hyp 0 prf) end;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   401
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   402
fun forall_intr_proof thy T (a, x) prf =
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   403
  let
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   404
    val Z = ztyp_of T;
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   405
    val z = global_zterm_of thy x;
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   406
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   407
    fun abs_term i b =
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   408
      if aconv_zterm (b, z) then ZBound i
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   409
      else
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   410
        (case b of
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   411
          ZAbs (x, T, t) => ZAbs (x, T, abs_term (i + 1) t)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   412
        | ZApp (t, u) => ZApp (abs_term i t, abs_term i u)
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   413
        | _ => b);
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   414
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   415
    fun abs_proof i (ZAbst (x, T, prf)) = ZAbst (x, T, abs_proof (i + 1) prf)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   416
      | abs_proof i (ZAbsP (x, t, prf)) = ZAbsP (x, abs_term i t, abs_proof i prf)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   417
      | abs_proof i (ZAppt (p, t)) = ZAppt (abs_proof i p, abs_term i t)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   418
      | abs_proof i (ZAppP (p, q)) = ZAppP (abs_proof i p, abs_proof i q)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   419
      | abs_proof _ p = p;
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   420
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   421
  in ZAbst (a, Z, abs_proof 0 prf) end;
79119
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   422
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   423
fun forall_elim_proof thy t p = ZAppt (p, global_zterm_of thy t);
cf29db6c95e1 more zterm operations;
wenzelm
parents: 79114
diff changeset
   424
79128
b6f5d4392388 more zproofs;
wenzelm
parents: 79126
diff changeset
   425
fun of_class_proof (T, c) = ZClassP (ztyp_of T, c);
b6f5d4392388 more zproofs;
wenzelm
parents: 79126
diff changeset
   426
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   427
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   428
(* equality *)
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   429
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   430
local
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   431
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   432
val thy0 =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   433
  Context.the_global_context ()
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   434
  |> Sign.add_types_global [(Binding.name "fun", 2, NoSyn), (Binding.name "prop", 0, NoSyn)]
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   435
  |> Sign.local_path
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   436
  |> Sign.add_consts
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   437
   [(Binding.name "all", (Term.aT [] --> propT) --> propT, NoSyn),
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   438
    (Binding.name "imp", propT --> propT --> propT, NoSyn),
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   439
    (Binding.name "eq", Term.aT [] --> Term.aT [] --> propT, NoSyn)];
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   440
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   441
val [reflexive_axiom, symmetric_axiom, transitive_axiom, equal_intr_axiom, equal_elim_axiom,
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   442
  abstract_rule_axiom, combination_axiom] =
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   443
    Theory.equality_axioms |> map (fn (b, t) => axiom_proof thy0 (Sign.full_name thy0 b) t);
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   444
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   445
in
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   446
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   447
val is_reflexive_proof =
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   448
  fn ZConstP (ZAxiom "Pure.reflexive", _, _, _) => true | _ => false;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   449
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   450
fun reflexive_proof thy T t =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   451
  let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   452
    val A = ztyp_of T;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   453
    val x = global_zterm_of thy t;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   454
  in map_const_proof (fn "'a" => A, fn "x" => x) reflexive_axiom end;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   455
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   456
fun symmetric_proof thy T t u prf =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   457
  if is_reflexive_proof prf then prf
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   458
  else
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   459
    let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   460
      val A = ztyp_of T;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   461
      val x = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   462
      val y = global_zterm_of thy u;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   463
      val ax = map_const_proof (fn "'a" => A, fn "x" => x | "y" => y) symmetric_axiom;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   464
    in ZAppP (ax, prf) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   465
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   466
fun transitive_proof thy T t u v prf1 prf2 =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   467
  if is_reflexive_proof prf1 then prf2
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   468
  else if is_reflexive_proof prf2 then prf1
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   469
  else
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   470
    let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   471
      val A = ztyp_of T;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   472
      val x = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   473
      val y = global_zterm_of thy u;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   474
      val z = global_zterm_of thy v;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   475
      val ax = map_const_proof (fn "'a" => A, fn "x" => x | "y" => y | "z" => z) transitive_axiom;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   476
    in ZAppP (ZAppP (ax, prf1), prf2) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   477
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   478
fun equal_intr_proof thy t u prf1 prf2 =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   479
  let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   480
    val A = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   481
    val B = global_zterm_of thy u;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   482
    val ax = map_const_proof (undefined, fn "A" => A | "B" => B) equal_intr_axiom;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   483
  in ZAppP (ZAppP (ax, prf1), prf2) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   484
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   485
fun equal_elim_proof thy t u prf1 prf2 =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   486
  let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   487
    val A = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   488
    val B = global_zterm_of thy u;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   489
    val ax = map_const_proof (undefined, fn "A" => A | "B" => B) equal_elim_axiom;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   490
  in ZAppP (ZAppP (ax, prf1), prf2) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   491
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   492
fun abstract_rule_proof thy T U x t u prf =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   493
  let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   494
    val A = ztyp_of T;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   495
    val B = ztyp_of U;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   496
    val f = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   497
    val g = global_zterm_of thy u;
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   498
    val ax =
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   499
      map_const_proof (fn "'a" => A | "'b" => B, fn "f" => f | "g" => g)
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   500
        abstract_rule_axiom;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   501
  in ZAppP (ax, forall_intr_proof thy T x prf) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   502
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   503
fun combination_proof thy T U f g t u prf1 prf2 =
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   504
  let
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   505
    val A = ztyp_of T;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   506
    val B = ztyp_of U;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   507
    val f' = global_zterm_of thy f;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   508
    val g' = global_zterm_of thy g;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   509
    val x = global_zterm_of thy t;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   510
    val y = global_zterm_of thy u;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   511
    val ax =
79126
bdb33a2d4167 clarified const_proof vs. zproof_name;
wenzelm
parents: 79124
diff changeset
   512
      map_const_proof (fn "'a" => A | "'b" => B, fn "f" => f' | "g" => g' | "x" => x | "y" => y)
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   513
        combination_axiom;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   514
  in ZAppP (ZAppP (ax, prf1), prf2) end;
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   515
79098
d8940e5bbb25 tight representation of types / terms / proof terms (presently unused);
wenzelm
parents:
diff changeset
   516
end;
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   517
79133
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   518
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   519
(* substitution *)
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   520
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   521
fun generalize_proof (tfrees, frees) idx prf =
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   522
  let
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   523
    val typ =
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   524
      if Names.is_empty tfrees then Same.same else
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   525
        subst_type_same (fn ((a, i), S) =>
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   526
          if i = ~1 andalso Names.defined tfrees a then ZTVar ((a, idx), S)
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   527
          else raise Same.SAME);
79136
wenzelm
parents: 79135
diff changeset
   528
    val term =
79147
bfe5c20074e4 proper substitution of types within term;
wenzelm
parents: 79146
diff changeset
   529
      subst_term_same typ (fn ((x, i), T) =>
bfe5c20074e4 proper substitution of types within term;
wenzelm
parents: 79146
diff changeset
   530
        if i = ~1 andalso Names.defined frees x then ZVar ((x, idx), T)
bfe5c20074e4 proper substitution of types within term;
wenzelm
parents: 79146
diff changeset
   531
        else raise Same.SAME);
79136
wenzelm
parents: 79135
diff changeset
   532
  in Same.commit (map_proof_same typ term) prf end;
79133
cfe995369655 more zproofs;
wenzelm
parents: 79132
diff changeset
   533
79149
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   534
fun instantiate_proof thy (Ts, ts) prf =
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   535
  let
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   536
    val instT = ZTVars.build (Ts |> fold (fn (v, T) => ZTVars.add (v, ztyp_of T)));
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   537
    val inst =
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   538
      ZVars.build (ts |> fold (fn ((v, T), t) => ZVars.add ((v, ztyp_of T), global_zterm_of thy t)));
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   539
    val typ =
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   540
      if ZTVars.is_empty instT then Same.same
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   541
      else subst_type_same (Same.function (ZTVars.lookup instT));
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   542
    val term = subst_term_same typ (Same.function (ZVars.lookup inst));
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   543
  in Same.commit (map_proof_same typ term) prf end;
810679c5ed3c more zproofs;
wenzelm
parents: 79148
diff changeset
   544
79135
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   545
fun varifyT_proof names prf =
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   546
  if null names then prf
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   547
  else
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   548
    let
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   549
      val tab = ZTVars.build (names |> fold (fn ((a, S), b) => ZTVars.add (((a, ~1), S), b)));
79136
wenzelm
parents: 79135
diff changeset
   550
      val typ =
wenzelm
parents: 79135
diff changeset
   551
        subst_type_same (fn v =>
wenzelm
parents: 79135
diff changeset
   552
          (case ZTVars.lookup tab v of
wenzelm
parents: 79135
diff changeset
   553
            NONE => raise Same.SAME
wenzelm
parents: 79135
diff changeset
   554
          | SOME w => ZTVar w));
79144
wenzelm
parents: 79136
diff changeset
   555
    in Same.commit (map_proof_types_same typ) prf end;
79135
db2dc7634d62 more zproofs;
wenzelm
parents: 79133
diff changeset
   556
79152
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   557
fun legacy_freezeT_proof t prf =
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   558
  (case Type.legacy_freezeT t of
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   559
    NONE => prf
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   560
  | SOME f =>
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   561
      let
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   562
        val tvar = ztyp_of o Same.function f;
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   563
        val typ = subst_type_same tvar;
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   564
      in Same.commit (map_proof_types_same typ) prf end);
4189e10f1524 more zproofs;
wenzelm
parents: 79150
diff changeset
   565
79150
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   566
fun incr_indexes_proof inc prf =
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   567
  let
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   568
    fun tvar ((a, i), S) = if i >= 0 then ZTVar ((a, i + inc), S) else raise Same.SAME;
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   569
    fun var ((x, i), T) = if i >= 0 then ZVar ((x, i + inc), T) else raise Same.SAME;
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   570
    val typ = subst_type_same tvar;
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   571
    val term = subst_term_same typ var;
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   572
  in Same.commit (map_proof_same typ term) prf end;
1cdc685fe852 more zproofs;
wenzelm
parents: 79149
diff changeset
   573
79124
89d4a8f52738 more zproofs;
wenzelm
parents: 79119
diff changeset
   574
end;