src/Pure/drule.ML
author paulson
Fri, 16 Feb 1996 12:08:49 +0100
changeset 1499 01fdd1ea6324
parent 1460 5a6f2aabd538
child 1596 4a09f4698813
permissions -rw-r--r--
Elimination of fully-functorial style. Type tactic changed to a type abbrevation (from a datatype). Constructor tactic and function apply deleted.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
     1
(*  Title:      Pure/drule.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Derived rules and other operations on theorems and theories
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
     9
infix 0 RS RSN RL RLN MRS MRL COMP;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
signature DRULE =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
  sig
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    13
  val add_defs		: (string * string) list -> theory -> theory
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    14
  val add_defs_i	: (string * term) list -> theory -> theory
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    15
  val asm_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    16
  val assume_ax		: theory -> string -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    17
  val COMP		: thm * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    18
  val compose		: thm * int * thm -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    19
  val cprems_of		: thm -> cterm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    20
  val cskip_flexpairs	: cterm -> cterm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    21
  val cstrip_imp_prems	: cterm -> cterm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    22
  val cterm_instantiate	: (cterm*cterm)list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    23
  val cut_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    24
  val equal_abs_elim	: cterm  -> thm -> thm
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    25
  val equal_abs_elim_list: cterm list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    26
  val eq_thm		: thm * thm -> bool
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    27
  val same_thm		: thm * thm -> bool
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    28
  val eq_thm_sg		: thm * thm -> bool
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    29
  val flexpair_abs_elim_list: cterm list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    30
  val forall_intr_list	: cterm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    31
  val forall_intr_frees	: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    32
  val forall_intr_vars	: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    33
  val forall_elim_list	: cterm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    34
  val forall_elim_var	: int -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    35
  val forall_elim_vars	: int -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    36
  val implies_elim_list	: thm -> thm list -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    37
  val implies_intr_list	: cterm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    38
  val MRL		: thm list list * thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    39
  val MRS		: thm list * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    40
  val pprint_cterm	: cterm -> pprint_args -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    41
  val pprint_ctyp	: ctyp -> pprint_args -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    42
  val pprint_theory	: theory -> pprint_args -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    43
  val pprint_thm	: thm -> pprint_args -> unit
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
    44
  val pretty_thm	: thm -> Pretty.T
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    45
  val print_cterm	: cterm -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    46
  val print_ctyp	: ctyp -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    47
  val print_goals	: int -> thm -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    48
  val print_goals_ref	: (int -> thm -> unit) ref
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    49
  val print_syntax	: theory -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    50
  val print_theory	: theory -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    51
  val print_thm		: thm -> unit
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    52
  val prth		: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    53
  val prthq		: thm Sequence.seq -> thm Sequence.seq
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    54
  val prths		: thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    55
  val read_instantiate	: (string*string)list -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
  val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    57
  val read_insts	:
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    58
          Sign.sg -> (indexname -> typ option) * (indexname -> sort option)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    59
                  -> (indexname -> typ option) * (indexname -> sort option)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
    60
                  -> string list -> (string*string)list
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    61
                  -> (indexname*ctyp)list * (cterm*cterm)list
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    62
  val reflexive_thm	: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    63
  val revcut_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    64
  val rewrite_goal_rule	: bool*bool -> (meta_simpset -> thm -> thm option)
214
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
    65
        -> meta_simpset -> int -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
  val rewrite_goals_rule: thm list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    67
  val rewrite_rule	: thm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    68
  val RS		: thm * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    69
  val RSN		: thm * (int * thm) -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    70
  val RL		: thm list * thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    71
  val RLN		: thm list * (int * thm list) -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    72
  val show_hyps		: bool ref
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    73
  val size_of_thm	: thm -> int
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    74
  val standard		: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    75
  val string_of_cterm	: cterm -> string
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    76
  val string_of_ctyp	: ctyp -> string
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    77
  val string_of_thm	: thm -> string
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    78
  val symmetric_thm	: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    79
  val thin_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    80
  val transitive_thm	: thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
  val triv_forall_equality: thm
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
  val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    83
  val zero_var_indexes	: thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
    86
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
    87
structure Drule : DRULE =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    90
(**** Extend Theories ****)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    91
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    92
(** add constant definitions **)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    93
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    94
(* all_axioms_of *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    95
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    96
(*results may contain duplicates!*)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    97
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    98
fun ancestry_of thy =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    99
  thy :: flat (map ancestry_of (parents_of thy));
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   100
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   101
val all_axioms_of =
776
df8f91c0e57c improved axioms_of: returns thms as the manual says;
wenzelm
parents: 708
diff changeset
   102
  flat o map (Symtab.dest o #new_axioms o rep_theory) o ancestry_of;
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   103
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   104
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   105
(* clash_types, clash_consts *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   106
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   107
(*check if types have common instance (ignoring sorts)*)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   108
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   109
fun clash_types ty1 ty2 =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   110
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   111
    val ty1' = Type.varifyT ty1;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   112
    val ty2' = incr_tvar (maxidx_of_typ ty1' + 1) (Type.varifyT ty2);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   113
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   114
    Type.raw_unify (ty1', ty2')
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   115
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   116
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   117
fun clash_consts (c1, ty1) (c2, ty2) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   118
  c1 = c2 andalso clash_types ty1 ty2;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   119
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   120
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   121
(* clash_defns *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   122
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   123
fun clash_defn c_ty (name, tm) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   124
  let val (c, ty') = dest_Const (head_of (fst (Logic.dest_equals tm))) in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   125
    if clash_consts c_ty (c, ty') then Some (name, ty') else None
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   126
  end handle TERM _ => None;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   127
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   128
fun clash_defns c_ty axms =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   129
  distinct (mapfilter (clash_defn c_ty) axms);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   130
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   131
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   132
(* dest_defn *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   133
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   134
fun dest_defn tm =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   135
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   136
    fun err msg = raise_term msg [tm];
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   137
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   138
    val (lhs, rhs) = Logic.dest_equals tm
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   139
      handle TERM _ => err "Not a meta-equality (==)";
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   140
    val (head, args) = strip_comb lhs;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   141
    val (c, ty) = dest_Const head
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   142
      handle TERM _ => err "Head of lhs not a constant";
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   143
655
9748dbcd4157 minor change of occs_const in dest_defn;
wenzelm
parents: 641
diff changeset
   144
    fun occs_const (Const c_ty') = (c_ty' = (c, ty))
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   145
      | occs_const (Abs (_, _, t)) = occs_const t
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   146
      | occs_const (t $ u) = occs_const t orelse occs_const u
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   147
      | occs_const _ = false;
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   148
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   149
    val show_frees = commas_quote o map (fst o dest_Free);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   150
    val show_tfrees = commas_quote o map fst;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   151
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   152
    val lhs_dups = duplicates args;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   153
    val rhs_extras = gen_rems (op =) (term_frees rhs, args);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   154
    val rhs_extrasT = gen_rems (op =) (term_tfrees rhs, typ_tfrees ty);
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   155
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   156
    if not (forall is_Free args) then
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   157
      err "Arguments of lhs have to be variables"
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   158
    else if not (null lhs_dups) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   159
      err ("Duplicate variables on lhs: " ^ show_frees lhs_dups)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   160
    else if not (null rhs_extras) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   161
      err ("Extra variables on rhs: " ^ show_frees rhs_extras)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   162
    else if not (null rhs_extrasT) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   163
      err ("Extra type variables on rhs: " ^ show_tfrees rhs_extrasT)
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   164
    else if occs_const rhs then
655
9748dbcd4157 minor change of occs_const in dest_defn;
wenzelm
parents: 641
diff changeset
   165
      err ("Constant to be defined occurs on rhs")
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   166
    else (c, ty)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   167
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   168
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   169
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   170
(* check_defn *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   171
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   172
fun err_in_defn name msg =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   173
  (writeln msg; error ("The error(s) above occurred in definition " ^ quote name));
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   174
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   175
fun check_defn sign (axms, (name, tm)) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   176
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   177
    fun show_const (c, ty) = quote (Pretty.string_of (Pretty.block
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   178
      [Pretty.str (c ^ " ::"), Pretty.brk 1, Sign.pretty_typ sign ty]));
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   179
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   180
    fun show_defn c (dfn, ty') = show_const (c, ty') ^ " in " ^ dfn;
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   181
    fun show_defns c = cat_lines o map (show_defn c);
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   182
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   183
    val (c, ty) = dest_defn tm
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   184
      handle TERM (msg, _) => err_in_defn name msg;
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   185
    val defns = clash_defns (c, ty) axms;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   186
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   187
    if not (null defns) then
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   188
      err_in_defn name ("Definition of " ^ show_const (c, ty) ^
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   189
        "\nclashes with " ^ show_defns c defns)
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   190
    else (name, tm) :: axms
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   191
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   192
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   193
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   194
(* add_defs *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   195
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   196
fun ext_defns prep_axm raw_axms thy =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   197
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   198
    val axms = map (prep_axm (sign_of thy)) raw_axms;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   199
    val all_axms = all_axioms_of thy;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   200
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   201
    foldl (check_defn (sign_of thy)) (all_axms, axms);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   202
    add_axioms_i axms thy
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   203
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   204
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   205
val add_defs_i = ext_defns cert_axm;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   206
val add_defs = ext_defns read_axm;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   207
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   208
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   209
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
(**** More derived rules and operations on theorems ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   211
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   212
(** some cterm->cterm operations: much faster than calling cterm_of! **)
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   213
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   214
(*Discard flexflex pairs; return a cterm*)
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   215
fun cskip_flexpairs ct =
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   216
    case term_of ct of
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   217
	(Const("==>", _) $ (Const("=?=",_)$_$_) $ _) =>
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   218
	    cskip_flexpairs (#2 (dest_cimplies ct))
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   219
      | _ => ct;
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   220
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   221
(* A1==>...An==>B  goes to  [A1,...,An], where B is not an implication *)
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   222
fun cstrip_imp_prems ct =
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   223
    let val (cA,cB) = dest_cimplies ct
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   224
    in  cA :: cstrip_imp_prems cB  end
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   225
    handle TERM _ => [];
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   226
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   227
(*The premises of a theorem, as a cterm list*)
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   228
val cprems_of = cstrip_imp_prems o cskip_flexpairs o cprop_of;
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   229
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   230
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   231
(** reading of instantiations **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   232
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   233
fun indexname cs = case Syntax.scan_varname cs of (v,[]) => v
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   234
        | _ => error("Lexical error in variable name " ^ quote (implode cs));
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   235
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   236
fun absent ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   237
  error("No such variable in term: " ^ Syntax.string_of_vname ixn);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   238
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   239
fun inst_failure ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   240
  error("Instantiation of " ^ Syntax.string_of_vname ixn ^ " fails");
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   241
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   242
(* this code is a bit of a mess. add_cterm could be simplified greatly if
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   243
   simultaneous instantiations were read or at least type checked
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   244
   simultaneously rather than one after the other. This would make the tricky
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   245
   composition of implicit type instantiations (parameter tye) superfluous.
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   246
*)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   247
fun read_insts sign (rtypes,rsorts) (types,sorts) used insts =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   248
let val {tsig,...} = Sign.rep_sg sign
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   249
    fun split([],tvs,vs) = (tvs,vs)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   250
      | split((sv,st)::l,tvs,vs) = (case explode sv of
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   251
                  "'"::cs => split(l,(indexname cs,st)::tvs,vs)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   252
                | cs => split(l,tvs,(indexname cs,st)::vs));
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   253
    val (tvs,vs) = split(insts,[],[]);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   254
    fun readT((a,i),st) =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   255
        let val ixn = ("'" ^ a,i);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   256
            val S = case rsorts ixn of Some S => S | None => absent ixn;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   257
            val T = Sign.read_typ (sign,sorts) st;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   258
        in if Type.typ_instance(tsig,T,TVar(ixn,S)) then (ixn,T)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   259
           else inst_failure ixn
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   260
        end
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   261
    val tye = map readT tvs;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   262
    fun add_cterm ((cts,tye,used), (ixn,st)) =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   263
        let val T = case rtypes ixn of
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   264
                      Some T => typ_subst_TVars tye T
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   265
                    | None => absent ixn;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   266
            val (ct,tye2) = read_def_cterm(sign,types,sorts) used false (st,T);
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   267
            val cts' = (ixn,T,ct)::cts
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   268
            fun inst(ixn,T,ct) = (ixn,typ_subst_TVars tye2 T,ct)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   269
            val used' = add_term_tvarnames(term_of ct,used);
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   270
        in (map inst cts',tye2 @ tye,used') end
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   271
    val (cterms,tye',_) = foldl add_cterm (([],tye,used), vs);
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   272
in (map (fn (ixn,T) => (ixn,ctyp_of sign T)) tye',
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   273
    map (fn (ixn,T,ct) => (cterm_of sign (Var(ixn,T)), ct)) cterms)
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   274
end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   275
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   276
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   277
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   278
(*** Printing of theories, theorems, etc. ***)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   279
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   280
(*If false, hypotheses are printed as dots*)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   281
val show_hyps = ref true;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   282
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   283
fun pretty_thm th =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   284
  let
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   285
    val {sign, hyps, prop, ...} = rep_thm th;
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   286
    val xshyps = extra_shyps th;
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   287
    val hlen = length xshyps + length hyps;
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   288
    val hsymbs =
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   289
      if hlen = 0 then []
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   290
      else if ! show_hyps then
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   291
        [Pretty.brk 2, Pretty.list "[" "]"
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   292
          (map (Sign.pretty_term sign) hyps @ map Sign.pretty_sort xshyps)]
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   293
      else
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   294
        [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ "]")];
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   295
  in
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   296
    Pretty.block (Sign.pretty_term sign prop :: hsymbs)
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   297
  end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   298
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   299
val string_of_thm = Pretty.string_of o pretty_thm;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   300
val pprint_thm = Pretty.pprint o Pretty.quote o pretty_thm;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   301
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   302
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   303
(** Top-level commands for printing theorems **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   304
val print_thm = writeln o string_of_thm;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   305
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   306
fun prth th = (print_thm th; th);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   307
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   308
(*Print and return a sequence of theorems, separated by blank lines. *)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   309
fun prthq thseq =
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   310
  (Sequence.prints (fn _ => print_thm) 100000 thseq; thseq);
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   311
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   312
(*Print and return a list of theorems, separated by blank lines. *)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   313
fun prths ths = (print_list_ln print_thm ths; ths);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   314
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   315
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   316
(* other printing commands *)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   317
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   318
fun pprint_ctyp cT =
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   319
  let val {sign, T} = rep_ctyp cT in Sign.pprint_typ sign T end;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   320
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   321
fun string_of_ctyp cT =
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   322
  let val {sign, T} = rep_ctyp cT in Sign.string_of_typ sign T end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   323
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   324
val print_ctyp = writeln o string_of_ctyp;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   325
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   326
fun pprint_cterm ct =
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   327
  let val {sign, t, ...} = rep_cterm ct in Sign.pprint_term sign t end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   328
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   329
fun string_of_cterm ct =
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   330
  let val {sign, t, ...} = rep_cterm ct in Sign.string_of_term sign t end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   331
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   332
val print_cterm = writeln o string_of_cterm;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   333
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   334
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   335
(* print theory *)
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   336
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   337
val pprint_theory = Sign.pprint_sg o sign_of;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   338
575
74f0e5fce609 added print_syntax: theory -> unit;
wenzelm
parents: 561
diff changeset
   339
val print_syntax = Syntax.print_syntax o syn_of;
74f0e5fce609 added print_syntax: theory -> unit;
wenzelm
parents: 561
diff changeset
   340
385
921f87897a76 added print_sign, print_axioms: theory -> unit;
wenzelm
parents: 252
diff changeset
   341
fun print_axioms thy =
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   342
  let
400
3c2c40c87112 replaced ext_axtab by new_axioms;
wenzelm
parents: 385
diff changeset
   343
    val {sign, new_axioms, ...} = rep_theory thy;
3c2c40c87112 replaced ext_axtab by new_axioms;
wenzelm
parents: 385
diff changeset
   344
    val axioms = Symtab.dest new_axioms;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   345
385
921f87897a76 added print_sign, print_axioms: theory -> unit;
wenzelm
parents: 252
diff changeset
   346
    fun prt_axm (a, t) = Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1,
921f87897a76 added print_sign, print_axioms: theory -> unit;
wenzelm
parents: 252
diff changeset
   347
      Pretty.quote (Sign.pretty_term sign t)];
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   348
  in
385
921f87897a76 added print_sign, print_axioms: theory -> unit;
wenzelm
parents: 252
diff changeset
   349
    Pretty.writeln (Pretty.big_list "additional axioms:" (map prt_axm axioms))
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   350
  end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   351
843
c1a4a4206102 removed print_sign, print_axioms;
wenzelm
parents: 776
diff changeset
   352
fun print_theory thy = (Sign.print_sg (sign_of thy); print_axioms thy);
385
921f87897a76 added print_sign, print_axioms: theory -> unit;
wenzelm
parents: 252
diff changeset
   353
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   354
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   355
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   356
(** Print thm A1,...,An/B in "goal style" -- premises as numbered subgoals **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   357
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   358
(* get type_env, sort_env of term *)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   359
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   360
local
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   361
  open Syntax;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   362
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   363
  fun ins_entry (x, y) [] = [(x, [y])]
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   364
    | ins_entry (x, y) ((pair as (x', ys')) :: pairs) =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   365
        if x = x' then (x', y ins ys') :: pairs
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   366
        else pair :: ins_entry (x, y) pairs;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   367
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   368
  fun add_type_env (Free (x, T), env) = ins_entry (T, x) env
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   369
    | add_type_env (Var (xi, T), env) = ins_entry (T, string_of_vname xi) env
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   370
    | add_type_env (Abs (_, _, t), env) = add_type_env (t, env)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   371
    | add_type_env (t $ u, env) = add_type_env (u, add_type_env (t, env))
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   372
    | add_type_env (_, env) = env;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   373
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   374
  fun add_sort_env (Type (_, Ts), env) = foldr add_sort_env (Ts, env)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   375
    | add_sort_env (TFree (x, S), env) = ins_entry (S, x) env
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   376
    | add_sort_env (TVar (xi, S), env) = ins_entry (S, string_of_vname xi) env;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   377
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   378
  val sort = map (apsnd sort_strings);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   379
in
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   380
  fun type_env t = sort (add_type_env (t, []));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   381
  fun sort_env t = rev (sort (it_term_types add_sort_env (t, [])));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   382
end;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   383
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   384
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   385
(* print_goals *)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   386
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   387
fun print_goals maxgoals state =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   388
  let
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   389
    open Syntax;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   390
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   391
    val {sign, prop, ...} = rep_thm state;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   392
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   393
    val pretty_term = Sign.pretty_term sign;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   394
    val pretty_typ = Sign.pretty_typ sign;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   395
    val pretty_sort = Sign.pretty_sort;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   396
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   397
    fun pretty_vars prtf (X, vs) = Pretty.block
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   398
      [Pretty.block (Pretty.commas (map Pretty.str vs)),
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   399
        Pretty.str " ::", Pretty.brk 1, prtf X];
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   400
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   401
    fun print_list _ _ [] = ()
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   402
      | print_list name prtf lst =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   403
          (writeln ""; Pretty.writeln (Pretty.big_list name (map prtf lst)));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   404
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   405
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   406
    fun print_goals (_, []) = ()
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   407
      | print_goals (n, A :: As) = (Pretty.writeln (Pretty.blk (0,
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   408
          [Pretty.str (" " ^ string_of_int n ^ ". "), pretty_term A]));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   409
            print_goals (n + 1, As));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   410
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   411
    val print_ffpairs =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   412
      print_list "Flex-flex pairs:" (pretty_term o Logic.mk_flexpair);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   413
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   414
    val print_types = print_list "Types:" (pretty_vars pretty_typ) o type_env;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   415
    val print_sorts = print_list "Sorts:" (pretty_vars pretty_sort) o sort_env;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   416
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   417
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   418
    val (tpairs, As, B) = Logic.strip_horn prop;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   419
    val ngoals = length As;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   420
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   421
    val orig_no_freeTs = ! show_no_free_types;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   422
    val orig_sorts = ! show_sorts;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   423
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   424
    fun restore () =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   425
      (show_no_free_types := orig_no_freeTs; show_sorts := orig_sorts);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   426
  in
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   427
    (show_no_free_types := true; show_sorts := false;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   428
      Pretty.writeln (pretty_term B);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   429
      if ngoals = 0 then writeln "No subgoals!"
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   430
      else if ngoals > maxgoals then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   431
        (print_goals (1, take (maxgoals, As));
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   432
          writeln ("A total of " ^ string_of_int ngoals ^ " subgoals..."))
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   433
      else print_goals (1, As);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   434
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   435
      print_ffpairs tpairs;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   436
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   437
      if orig_sorts then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   438
        (print_types prop; print_sorts prop)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   439
      else if ! show_types then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   440
        print_types prop
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   441
      else ())
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   442
    handle exn => (restore (); raise exn);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   443
    restore ()
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   444
  end;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   445
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   446
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   447
(*"hook" for user interfaces: allows print_goals to be replaced*)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   448
val print_goals_ref = ref print_goals;
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   449
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   450
(*** Find the type (sort) associated with a (T)Var or (T)Free in a term
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   451
     Used for establishing default types (of variables) and sorts (of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   452
     type variables) when reading another term.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   453
     Index -1 indicates that a (T)Free rather than a (T)Var is wanted.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   454
***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   455
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   456
fun types_sorts thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   457
    let val {prop,hyps,...} = rep_thm thm;
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   458
        val big = list_comb(prop,hyps); (* bogus term! *)
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   459
        val vars = map dest_Var (term_vars big);
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   460
        val frees = map dest_Free (term_frees big);
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   461
        val tvars = term_tvars big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   462
        val tfrees = term_tfrees big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   463
        fun typ(a,i) = if i<0 then assoc(frees,a) else assoc(vars,(a,i));
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   464
        fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   465
    in (typ,sort) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   466
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   467
(** Standardization of rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   468
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   469
(*Generalization over a list of variables, IGNORING bad ones*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   470
fun forall_intr_list [] th = th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   471
  | forall_intr_list (y::ys) th =
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   472
        let val gth = forall_intr_list ys th
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   473
        in  forall_intr y gth   handle THM _ =>  gth  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   475
(*Generalization over all suitable Free variables*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   476
fun forall_intr_frees th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   477
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   478
    in  forall_intr_list
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   479
         (map (cterm_of sign) (sort atless (term_frees prop)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   480
         th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   481
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   482
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   483
(*Replace outermost quantified variable by Var of given index.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   484
    Could clash with Vars already present.*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   485
fun forall_elim_var i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   486
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   487
    in case prop of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   488
          Const("all",_) $ Abs(a,T,_) =>
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   489
              forall_elim (cterm_of sign (Var((a,i), T)))  th
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   490
        | _ => raise THM("forall_elim_var", i, [th])
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   491
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   492
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   493
(*Repeat forall_elim_var until all outer quantifiers are removed*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   494
fun forall_elim_vars i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   495
    forall_elim_vars i (forall_elim_var i th)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   496
        handle THM _ => th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   497
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   498
(*Specialization over a list of cterms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   499
fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   500
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   501
(* maps [A1,...,An], B   to   [| A1;...;An |] ==> B  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   502
fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   503
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   504
(* maps [| A1;...;An |] ==> B and [A1,...,An]   to   B *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   505
fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   506
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   507
(*Reset Var indexes to zero, renaming to preserve distinctness*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   508
fun zero_var_indexes th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   509
    let val {prop,sign,...} = rep_thm th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   510
        val vars = term_vars prop
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   511
        val bs = foldl add_new_id ([], map (fn Var((a,_),_)=>a) vars)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   512
        val inrs = add_term_tvars(prop,[]);
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   513
        val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs));
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   514
        val tye = map (fn ((v,rs),a) => (v, TVar((a,0),rs))) (inrs ~~ nms')
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   515
        val ctye = map (fn (v,T) => (v,ctyp_of sign T)) tye;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   516
        fun varpairs([],[]) = []
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   517
          | varpairs((var as Var(v,T)) :: vars, b::bs) =
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   518
                let val T' = typ_subst_TVars tye T
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   519
                in (cterm_of sign (Var(v,T')),
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   520
                    cterm_of sign (Var((b,0),T'))) :: varpairs(vars,bs)
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   521
                end
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   522
          | varpairs _ = raise TERM("varpairs", []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   523
    in instantiate (ctye, varpairs(vars,rev bs)) th end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   524
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   525
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   526
(*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   527
    all generality expressed by Vars having index 0.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   528
fun standard th =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   529
  let val {maxidx,...} = rep_thm th
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   530
  in
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   531
    th |> implies_intr_hyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   532
       |> forall_intr_frees |> forall_elim_vars (maxidx + 1)
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   533
       |> Thm.strip_shyps |> Thm.implies_intr_shyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   534
       |> zero_var_indexes |> Thm.varifyT |> Thm.compress
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   535
  end;
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   536
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   537
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   538
(*Assume a new formula, read following the same conventions as axioms.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   539
  Generalizes over Free variables,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   540
  creates the assumption, and then strips quantifiers.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   541
  Example is [| ALL x:?A. ?P(x) |] ==> [| ?P(?a) |]
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   542
             [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ]    *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   543
fun assume_ax thy sP =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   544
    let val sign = sign_of thy
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   545
        val prop = Logic.close_form (term_of (read_cterm sign
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   546
                         (sP, propT)))
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   547
    in forall_elim_vars 0 (assume (cterm_of sign prop))  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   548
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   549
(*Resolution: exactly one resolvent must be produced.*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   550
fun tha RSN (i,thb) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   551
  case Sequence.chop (2, biresolution false [(false,tha)] i thb) of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   552
      ([th],_) => th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   553
    | ([],_)   => raise THM("RSN: no unifiers", i, [tha,thb])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   554
    |      _   => raise THM("RSN: multiple unifiers", i, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   555
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   556
(*resolution: P==>Q, Q==>R gives P==>R. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   557
fun tha RS thb = tha RSN (1,thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   558
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   559
(*For joining lists of rules*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   560
fun thas RLN (i,thbs) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   561
  let val resolve = biresolution false (map (pair false) thas) i
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   562
      fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => []
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   563
  in  flat (map resb thbs)  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   564
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   565
fun thas RL thbs = thas RLN (1,thbs);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   566
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   567
(*Resolve a list of rules against bottom_rl from right to left;
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   568
  makes proof trees*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   569
fun rls MRS bottom_rl =
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   570
  let fun rs_aux i [] = bottom_rl
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   571
        | rs_aux i (rl::rls) = rl RSN (i, rs_aux (i+1) rls)
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   572
  in  rs_aux 1 rls  end;
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   573
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   574
(*As above, but for rule lists*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   575
fun rlss MRL bottom_rls =
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   576
  let fun rs_aux i [] = bottom_rls
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   577
        | rs_aux i (rls::rlss) = rls RLN (i, rs_aux (i+1) rlss)
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   578
  in  rs_aux 1 rlss  end;
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   579
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   580
(*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   581
  with no lifting or renaming!  Q may contain ==> or meta-quants
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   582
  ALWAYS deletes premise i *)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   583
fun compose(tha,i,thb) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   584
    Sequence.list_of_s (bicompose false (false,tha,0) i thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   585
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   586
(*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   587
fun tha COMP thb =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   588
    case compose(tha,1,thb) of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   589
        [th] => th
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   590
      | _ =>   raise THM("COMP", 1, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   591
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   592
(*Instantiate theorem th, reading instantiations under signature sg*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   593
fun read_instantiate_sg sg sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   594
    let val ts = types_sorts th;
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   595
        val used = add_term_tvarnames(#prop(rep_thm th),[]);
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   596
    in  instantiate (read_insts sg ts ts used sinsts) th  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   597
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   598
(*Instantiate theorem th, reading instantiations under theory of th*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   599
fun read_instantiate sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   600
    read_instantiate_sg (#sign (rep_thm th)) sinsts th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   601
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   602
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   603
(*Left-to-right replacements: tpairs = [...,(vi,ti),...].
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   604
  Instantiates distinct Vars by terms, inferring type instantiations. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   605
local
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   606
  fun add_types ((ct,cu), (sign,tye,maxidx)) =
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   607
    let val {sign=signt, t=t, T= T, maxidx=maxidxt,...} = rep_cterm ct
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   608
        and {sign=signu, t=u, T= U, maxidx=maxidxu,...} = rep_cterm cu;
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   609
        val maxi = max[maxidx,maxidxt,maxidxu];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   610
        val sign' = Sign.merge(sign, Sign.merge(signt, signu))
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   611
        val (tye',maxi') = Type.unify (#tsig(Sign.rep_sg sign')) maxi tye (T,U)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   612
          handle Type.TUNIFY => raise TYPE("add_types", [T,U], [t,u])
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   613
    in  (sign', tye', maxi')  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   614
in
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   615
fun cterm_instantiate ctpairs0 th =
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   616
  let val (sign,tye,_) = foldr add_types (ctpairs0, (#sign(rep_thm th),[],0))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   617
      val tsig = #tsig(Sign.rep_sg sign);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   618
      fun instT(ct,cu) = let val inst = subst_TVars tye
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   619
                         in (cterm_fun inst ct, cterm_fun inst cu) end
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   620
      fun ctyp2 (ix,T) = (ix, ctyp_of sign T)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   621
  in  instantiate (map ctyp2 tye, map instT ctpairs0) th  end
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   622
  handle TERM _ =>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   623
           raise THM("cterm_instantiate: incompatible signatures",0,[th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   624
       | TYPE _ => raise THM("cterm_instantiate: types", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   625
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   626
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   627
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   628
(** theorem equality test is exported and used by BEST_FIRST **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   629
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   630
(*equality of theorems uses equality of signatures and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   631
  the a-convertible test for terms*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   632
fun eq_thm (th1,th2) =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   633
    let val {sign=sg1, shyps=shyps1, hyps=hyps1, prop=prop1, ...} = rep_thm th1
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   634
        and {sign=sg2, shyps=shyps2, hyps=hyps2, prop=prop2, ...} = rep_thm th2
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   635
    in  Sign.eq_sg (sg1,sg2) andalso
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   636
        eq_set (shyps1, shyps2) andalso
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   637
        aconvs(hyps1,hyps2) andalso
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   638
        prop1 aconv prop2
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   639
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   640
1241
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   641
(*equality of theorems using similarity of signatures,
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   642
  i.e. the theorems belong to the same theory but not necessarily to the same
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   643
  version of this theory*)
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   644
fun same_thm (th1,th2) =
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   645
    let val {sign=sg1, shyps=shyps1, hyps=hyps1, prop=prop1, ...} = rep_thm th1
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   646
        and {sign=sg2, shyps=shyps2, hyps=hyps2, prop=prop2, ...} = rep_thm th2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   647
    in  Sign.same_sg (sg1,sg2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   648
        eq_set (shyps1, shyps2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   649
        aconvs(hyps1,hyps2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   650
        prop1 aconv prop2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   651
    end;
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   652
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   653
(*Do the two theorems have the same signature?*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   654
fun eq_thm_sg (th1,th2) = Sign.eq_sg(#sign(rep_thm th1), #sign(rep_thm th2));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   655
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   656
(*Useful "distance" function for BEST_FIRST*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   657
val size_of_thm = size_of_term o #prop o rep_thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   658
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   659
1194
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   660
(** Mark Staples's weaker version of eq_thm: ignores variable renaming and
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   661
    (some) type variable renaming **)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   662
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   663
 (* Can't use term_vars, because it sorts the resulting list of variable names.
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   664
    We instead need the unique list noramlised by the order of appearance
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   665
    in the term. *)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   666
fun term_vars' (t as Var(v,T)) = [t]
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   667
  | term_vars' (Abs(_,_,b)) = term_vars' b
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   668
  | term_vars' (f$a) = (term_vars' f) @ (term_vars' a)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   669
  | term_vars' _ = [];
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   670
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   671
fun forall_intr_vars th =
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   672
  let val {prop,sign,...} = rep_thm th;
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   673
      val vars = distinct (term_vars' prop);
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   674
  in forall_intr_list (map (cterm_of sign) vars) th end;
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   675
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   676
fun weak_eq_thm (tha,thb) =
1194
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   677
    eq_thm(forall_intr_vars (freezeT tha), forall_intr_vars (freezeT thb));
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   678
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   679
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   680
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   681
(*** Meta-Rewriting Rules ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   682
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   683
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   684
val reflexive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   685
  let val cx = cterm_of Sign.proto_pure (Var(("x",0),TVar(("'a",0),logicS)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   686
  in Thm.reflexive cx end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   687
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   688
val symmetric_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   689
  let val xy = read_cterm Sign.proto_pure ("x::'a::logic == y",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   690
  in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   691
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   692
val transitive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   693
  let val xy = read_cterm Sign.proto_pure ("x::'a::logic == y",propT)
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   694
      val yz = read_cterm Sign.proto_pure ("y::'a::logic == z",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   695
      val xythm = Thm.assume xy and yzthm = Thm.assume yz
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   696
  in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   697
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   698
(** Below, a "conversion" has type cterm -> thm **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   699
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   700
val refl_cimplies = reflexive (cterm_of Sign.proto_pure implies);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   701
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   702
(*In [A1,...,An]==>B, rewrite the selected A's only -- for rewrite_goals_tac*)
214
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   703
(*Do not rewrite flex-flex pairs*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   704
fun goals_conv pred cv =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   705
  let fun gconv i ct =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   706
        let val (A,B) = Thm.dest_cimplies ct
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   707
            val (thA,j) = case term_of A of
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   708
                  Const("=?=",_)$_$_ => (reflexive A, i)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   709
                | _ => (if pred i then cv A else reflexive A, i+1)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   710
        in  combination (combination refl_cimplies thA) (gconv j B) end
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   711
        handle TERM _ => reflexive ct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   712
  in gconv 1 end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   713
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   714
(*Use a conversion to transform a theorem*)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   715
fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   716
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   717
(*rewriting conversion*)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   718
fun rew_conv mode prover mss = rewrite_cterm mode mss prover;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   719
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   720
(*Rewrite a theorem*)
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   721
fun rewrite_rule []   th = th
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   722
  | rewrite_rule thms th =
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   723
	fconv_rule (rew_conv (true,false) (K(K None)) (Thm.mss_of thms)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   724
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   725
(*Rewrite the subgoals of a proof state (represented by a theorem) *)
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   726
fun rewrite_goals_rule []   th = th
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   727
  | rewrite_goals_rule thms th =
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   728
	fconv_rule (goals_conv (K true) 
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   729
		    (rew_conv (true,false) (K(K None))
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   730
		     (Thm.mss_of thms))) 
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   731
	           th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   732
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   733
(*Rewrite the subgoal of a proof state (represented by a theorem) *)
214
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   734
fun rewrite_goal_rule mode prover mss i thm =
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   735
  if 0 < i  andalso  i <= nprems_of thm
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   736
  then fconv_rule (goals_conv (fn j => j=i) (rew_conv mode prover mss)) thm
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   737
  else raise THM("rewrite_goal_rule",i,[thm]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   738
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   739
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   740
(** Derived rules mainly for METAHYPS **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   741
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   742
(*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   743
fun equal_abs_elim ca eqth =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   744
  let val {sign=signa, t=a, ...} = rep_cterm ca
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   745
      and combth = combination eqth (reflexive ca)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   746
      val {sign,prop,...} = rep_thm eqth
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   747
      val (abst,absu) = Logic.dest_equals prop
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   748
      val cterm = cterm_of (Sign.merge (sign,signa))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   749
  in  transitive (symmetric (beta_conversion (cterm (abst$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   750
           (transitive combth (beta_conversion (cterm (absu$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   751
  end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   752
  handle THM _ => raise THM("equal_abs_elim", 0, [eqth]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   753
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   754
(*Calling equal_abs_elim with multiple terms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   755
fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   756
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   757
local
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   758
  open Logic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   759
  val alpha = TVar(("'a",0), [])     (*  type ?'a::{}  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   760
  fun err th = raise THM("flexpair_inst: ", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   761
  fun flexpair_inst def th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   762
    let val {prop = Const _ $ t $ u,  sign,...} = rep_thm th
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   763
        val cterm = cterm_of sign
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   764
        fun cvar a = cterm(Var((a,0),alpha))
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   765
        val def' = cterm_instantiate [(cvar"t", cterm t), (cvar"u", cterm u)]
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   766
                   def
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   767
    in  equal_elim def' th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   768
    end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   769
    handle THM _ => err th | bind => err th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   770
in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   771
val flexpair_intr = flexpair_inst (symmetric flexpair_def)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   772
and flexpair_elim = flexpair_inst flexpair_def
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   773
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   774
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   775
(*Version for flexflex pairs -- this supports lifting.*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   776
fun flexpair_abs_elim_list cts =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   777
    flexpair_intr o equal_abs_elim_list cts o flexpair_elim;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   778
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   779
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   780
(*** Some useful meta-theorems ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   781
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   782
(*The rule V/V, obtains assumption solving for eresolve_tac*)
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   783
val asm_rl = trivial(read_cterm Sign.proto_pure ("PROP ?psi",propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   784
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   785
(*Meta-level cut rule: [| V==>W; V |] ==> W *)
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   786
val cut_rl = trivial(read_cterm Sign.proto_pure
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   787
        ("PROP ?psi ==> PROP ?theta", propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   788
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   789
(*Generalized elim rule for one conclusion; cut_rl with reversed premises:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   790
     [| PROP V;  PROP V ==> PROP W |] ==> PROP W *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   791
val revcut_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   792
  let val V = read_cterm Sign.proto_pure ("PROP V", propT)
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   793
      and VW = read_cterm Sign.proto_pure ("PROP V ==> PROP W", propT);
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   794
  in  standard (implies_intr V
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   795
                (implies_intr VW
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   796
                 (implies_elim (assume VW) (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   797
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   798
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   799
(*for deleting an unwanted assumption*)
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   800
val thin_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   801
  let val V = read_cterm Sign.proto_pure ("PROP V", propT)
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   802
      and W = read_cterm Sign.proto_pure ("PROP W", propT);
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   803
  in  standard (implies_intr V (implies_intr W (assume W)))
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   804
  end;
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   805
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   806
(* (!!x. PROP ?V) == PROP ?V       Allows removal of redundant parameters*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   807
val triv_forall_equality =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   808
  let val V  = read_cterm Sign.proto_pure ("PROP V", propT)
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   809
      and QV = read_cterm Sign.proto_pure ("!!x::'a. PROP V", propT)
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   810
      and x  = read_cterm Sign.proto_pure ("x", TFree("'a",logicS));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   811
  in  standard (equal_intr (implies_intr QV (forall_elim x (assume QV)))
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   812
                           (implies_intr V  (forall_intr x (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   813
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   814
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   815
end;
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   816
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
   817
open Drule;