src/Pure/drule.ML
author paulson
Wed, 21 Aug 1996 13:22:23 +0200
changeset 1933 8b24773de6db
parent 1906 4699a9058a4f
child 2004 3411fe560611
permissions -rw-r--r--
Addition of message NS5
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
1703
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
    38
  val dest_cimplies     : cterm -> cterm * cterm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    39
  val MRL		: thm list list * thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    40
  val MRS		: thm list * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    41
  val read_instantiate	: (string*string)list -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
  val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    43
  val read_insts	:
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    44
          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
    45
                  -> (indexname -> typ option) * (indexname -> sort option)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
    46
                  -> string list -> (string*string)list
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    47
                  -> (indexname*ctyp)list * (cterm*cterm)list
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    48
  val reflexive_thm	: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    49
  val revcut_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    50
  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
    51
        -> meta_simpset -> int -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
  val rewrite_goals_rule: thm list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    53
  val rewrite_rule	: thm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    54
  val RS		: thm * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    55
  val RSN		: thm * (int * thm) -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    56
  val RL		: thm list * thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    57
  val RLN		: thm list * (int * thm list) -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    58
  val size_of_thm	: thm -> int
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    59
  val standard		: thm -> thm
1756
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
    60
  val swap_prems_rl     : thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    61
  val symmetric_thm	: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    62
  val thin_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    63
  val transitive_thm	: thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
  val triv_forall_equality: thm
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
  val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    66
  val zero_var_indexes	: thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
    69
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
    70
structure Drule : DRULE =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    73
(**** Extend Theories ****)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    74
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    75
(** add constant definitions **)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    76
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    77
(* all_axioms_of *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    78
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    79
(*results may contain duplicates!*)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    80
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    81
fun ancestry_of thy =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    82
  thy :: flat (map ancestry_of (parents_of thy));
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    83
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
    84
val all_axioms_of =
776
df8f91c0e57c improved axioms_of: returns thms as the manual says;
wenzelm
parents: 708
diff changeset
    85
  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
    86
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    87
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    88
(* clash_types, clash_consts *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    89
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    90
(*check if types have common instance (ignoring sorts)*)
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
fun clash_types ty1 ty2 =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    93
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    94
    val ty1' = Type.varifyT ty1;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    95
    val ty2' = incr_tvar (maxidx_of_typ ty1' + 1) (Type.varifyT ty2);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    96
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    97
    Type.raw_unify (ty1', ty2')
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    98
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
    99
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   100
fun clash_consts (c1, ty1) (c2, ty2) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   101
  c1 = c2 andalso clash_types ty1 ty2;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   102
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
(* clash_defns *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   105
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   106
fun clash_defn c_ty (name, tm) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   107
  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
   108
    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
   109
  end handle TERM _ => None;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   110
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   111
fun clash_defns c_ty axms =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   112
  distinct (mapfilter (clash_defn c_ty) axms);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   113
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   114
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   115
(* dest_defn *)
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 dest_defn tm =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   118
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   119
    fun err msg = raise_term msg [tm];
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
    val (lhs, rhs) = Logic.dest_equals tm
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   122
      handle TERM _ => err "Not a meta-equality (==)";
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   123
    val (head, args) = strip_comb lhs;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   124
    val (c, ty) = dest_Const head
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   125
      handle TERM _ => err "Head of lhs not a constant";
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   126
655
9748dbcd4157 minor change of occs_const in dest_defn;
wenzelm
parents: 641
diff changeset
   127
    fun occs_const (Const c_ty') = (c_ty' = (c, ty))
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   128
      | occs_const (Abs (_, _, t)) = occs_const t
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   129
      | occs_const (t $ u) = occs_const t orelse occs_const u
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   130
      | occs_const _ = false;
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   131
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   132
    val show_frees = commas_quote o map (fst o dest_Free);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   133
    val show_tfrees = commas_quote o map fst;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   134
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   135
    val lhs_dups = duplicates args;
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   136
    val rhs_extras = gen_rems (op =) (term_frees rhs, args);
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   137
    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
   138
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   139
    if not (forall is_Free args) then
1906
4699a9058a4f Improved (?) wording of error message
paulson
parents: 1756
diff changeset
   140
      err "Arguments (on lhs) must be variables"
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   141
    else if not (null lhs_dups) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   142
      err ("Duplicate variables on lhs: " ^ show_frees lhs_dups)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   143
    else if not (null rhs_extras) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   144
      err ("Extra variables on rhs: " ^ show_frees rhs_extras)
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   145
    else if not (null rhs_extrasT) then
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   146
      err ("Extra type variables on rhs: " ^ show_tfrees rhs_extrasT)
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   147
    else if occs_const rhs then
655
9748dbcd4157 minor change of occs_const in dest_defn;
wenzelm
parents: 641
diff changeset
   148
      err ("Constant to be defined occurs on rhs")
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   149
    else (c, ty)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   150
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   151
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   152
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   153
(* check_defn *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   154
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   155
fun err_in_defn name msg =
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   156
  (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
   157
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   158
fun check_defn sign (axms, (name, tm)) =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   159
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   160
    fun show_const (c, ty) = quote (Pretty.string_of (Pretty.block
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   161
      [Pretty.str (c ^ " ::"), Pretty.brk 1, Sign.pretty_typ sign ty]));
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   162
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   163
    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
   164
    fun show_defns c = cat_lines o map (show_defn c);
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   165
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   166
    val (c, ty) = dest_defn tm
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   167
      handle TERM (msg, _) => err_in_defn name msg;
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   168
    val defns = clash_defns (c, ty) axms;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   169
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   170
    if not (null defns) then
641
49fc43cd6a35 add_defs: improved error messages;
wenzelm
parents: 575
diff changeset
   171
      err_in_defn name ("Definition of " ^ show_const (c, ty) ^
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   172
        "\nclashes with " ^ show_defns c defns)
561
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   173
    else (name, tm) :: axms
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   174
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   175
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   176
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   177
(* add_defs *)
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   178
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   179
fun ext_defns prep_axm raw_axms thy =
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   180
  let
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   181
    val axms = map (prep_axm (sign_of thy)) raw_axms;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   182
    val all_axms = all_axioms_of thy;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   183
  in
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   184
    foldl (check_defn (sign_of thy)) (all_axms, axms);
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   185
    add_axioms_i axms thy
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   186
  end;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   187
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   188
val add_defs_i = ext_defns cert_axm;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   189
val add_defs = ext_defns read_axm;
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   190
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   191
95225e63ef02 added add_defs, add_defs_i;
wenzelm
parents: 400
diff changeset
   192
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
(**** More derived rules and operations on theorems ****)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   195
(** 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
   196
1703
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   197
(*dest_implies for cterms. Note T=prop below*)
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   198
fun dest_cimplies ct =
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   199
  (let val (ct1, ct2) = dest_comb ct
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   200
       val {t, ...} = rep_cterm ct1;
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   201
   in case head_of t of
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   202
          Const("==>", _) => (snd (dest_comb ct1), ct2)
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   203
        | _ => raise TERM ("dest_cimplies", [term_of ct])
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   204
   end
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   205
  ) handle CTERM "dest_comb" => raise TERM ("dest_cimplies", [term_of ct]);
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   206
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   207
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   208
(*Discard flexflex pairs; return a cterm*)
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   209
fun cskip_flexpairs ct =
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   210
    case term_of ct of
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   211
	(Const("==>", _) $ (Const("=?=",_)$_$_) $ _) =>
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   212
	    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
   213
      | _ => ct;
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   214
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   215
(* 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
   216
fun cstrip_imp_prems ct =
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   217
    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
   218
    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
   219
    handle TERM _ => [];
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
(*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
   222
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
   223
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   224
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   225
(** reading of instantiations **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   226
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   227
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
   228
        | _ => 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
   229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   230
fun absent ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   231
  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
   232
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   233
fun inst_failure ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   234
  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
   235
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   236
(* 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
   237
   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
   238
   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
   239
   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
   240
*)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   241
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
   242
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
   243
    fun split([],tvs,vs) = (tvs,vs)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   244
      | 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
   245
                  "'"::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
   246
                | 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
   247
    val (tvs,vs) = split(insts,[],[]);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   248
    fun readT((a,i),st) =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   249
        let val ixn = ("'" ^ a,i);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   250
            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
   251
            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
   252
        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
   253
           else inst_failure ixn
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   254
        end
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   255
    val tye = map readT tvs;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   256
    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
   257
        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
   258
                      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
   259
                    | None => absent ixn;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   260
            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
   261
            val cts' = (ixn,T,ct)::cts
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   262
            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
   263
            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
   264
        in (map inst cts',tye2 @ tye,used') end
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   265
    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
   266
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
   267
    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
   268
end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   269
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   270
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   271
(*** Find the type (sort) associated with a (T)Var or (T)Free in a term
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   272
     Used for establishing default types (of variables) and sorts (of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   273
     type variables) when reading another term.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
     Index -1 indicates that a (T)Free rather than a (T)Var is wanted.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   277
fun types_sorts thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   278
    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
   279
        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
   280
        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
   281
        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
   282
        val tvars = term_tvars big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   283
        val tfrees = term_tfrees big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   284
        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
   285
        fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
    in (typ,sort) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   287
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
(** Standardization of rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   289
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
(*Generalization over a list of variables, IGNORING bad ones*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
fun forall_intr_list [] th = th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
  | forall_intr_list (y::ys) th =
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   293
        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
   294
        in  forall_intr y gth   handle THM _ =>  gth  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   296
(*Generalization over all suitable Free variables*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   297
fun forall_intr_frees th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   299
    in  forall_intr_list
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   300
         (map (cterm_of sign) (sort atless (term_frees prop)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
         th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   303
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
(*Replace outermost quantified variable by Var of given index.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
    Could clash with Vars already present.*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   306
fun forall_elim_var i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
    in case prop of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   309
          Const("all",_) $ Abs(a,T,_) =>
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   310
              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
   311
        | _ => raise THM("forall_elim_var", i, [th])
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   314
(*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
   315
fun forall_elim_vars i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   316
    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
   317
        handle THM _ => th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   319
(*Specialization over a list of cterms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   320
fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   321
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   322
(* maps [A1,...,An], B   to   [| A1;...;An |] ==> B  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   324
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   325
(* maps [| A1;...;An |] ==> B and [A1,...,An]   to   B *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   327
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   328
(*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
   329
fun zero_var_indexes th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   330
    let val {prop,sign,...} = rep_thm th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   331
        val vars = term_vars prop
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   332
        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
   333
        val inrs = add_term_tvars(prop,[]);
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   334
        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
   335
        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
   336
        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
   337
        fun varpairs([],[]) = []
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   338
          | 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
   339
                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
   340
                in (cterm_of sign (Var(v,T')),
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   341
                    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
   342
                end
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   343
          | varpairs _ = raise TERM("varpairs", []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   344
    in instantiate (ctye, varpairs(vars,rev bs)) th end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   345
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   346
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   347
(*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   348
    all generality expressed by Vars having index 0.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   349
fun standard th =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   350
  let val {maxidx,...} = rep_thm th
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   351
  in
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   352
    th |> implies_intr_hyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   353
       |> forall_intr_frees |> forall_elim_vars (maxidx + 1)
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   354
       |> Thm.strip_shyps |> Thm.implies_intr_shyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   355
       |> zero_var_indexes |> Thm.varifyT |> Thm.compress
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   356
  end;
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   357
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   358
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   359
(*Assume a new formula, read following the same conventions as axioms.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   360
  Generalizes over Free variables,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   361
  creates the assumption, and then strips quantifiers.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   362
  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
   363
             [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ]    *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
fun assume_ax thy sP =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
    let val sign = sign_of thy
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   366
        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
   367
                         (sP, propT)))
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   368
    in forall_elim_vars 0 (assume (cterm_of sign prop))  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   369
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   370
(*Resolution: exactly one resolvent must be produced.*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
fun tha RSN (i,thb) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
  case Sequence.chop (2, biresolution false [(false,tha)] i thb) of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   373
      ([th],_) => th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   374
    | ([],_)   => raise THM("RSN: no unifiers", i, [tha,thb])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   375
    |      _   => raise THM("RSN: multiple unifiers", i, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   377
(*resolution: P==>Q, Q==>R gives P==>R. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   378
fun tha RS thb = tha RSN (1,thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   379
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   380
(*For joining lists of rules*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   381
fun thas RLN (i,thbs) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   382
  let val resolve = biresolution false (map (pair false) thas) i
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   383
      fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => []
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   384
  in  flat (map resb thbs)  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   385
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   386
fun thas RL thbs = thas RLN (1,thbs);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   387
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   388
(*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
   389
  makes proof trees*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   390
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
   391
  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
   392
        | 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
   393
  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
   394
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   395
(*As above, but for rule lists*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   396
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
   397
  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
   398
        | 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
   399
  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
   400
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   401
(*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   402
  with no lifting or renaming!  Q may contain ==> or meta-quants
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   403
  ALWAYS deletes premise i *)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   404
fun compose(tha,i,thb) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   405
    Sequence.list_of_s (bicompose false (false,tha,0) i thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   406
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   407
(*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
fun tha COMP thb =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   409
    case compose(tha,1,thb) of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   410
        [th] => th
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   411
      | _ =>   raise THM("COMP", 1, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   412
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   413
(*Instantiate theorem th, reading instantiations under signature sg*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   414
fun read_instantiate_sg sg sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   415
    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
   416
        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
   417
    in  instantiate (read_insts sg ts ts used sinsts) th  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   418
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   419
(*Instantiate theorem th, reading instantiations under theory of th*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   420
fun read_instantiate sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
    read_instantiate_sg (#sign (rep_thm th)) sinsts th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   422
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   424
(*Left-to-right replacements: tpairs = [...,(vi,ti),...].
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   425
  Instantiates distinct Vars by terms, inferring type instantiations. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   426
local
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   427
  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
   428
    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
   429
        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
   430
        val maxi = max[maxidx,maxidxt,maxidxu];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   431
        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
   432
        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
   433
          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
   434
    in  (sign', tye', maxi')  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   435
in
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   436
fun cterm_instantiate ctpairs0 th =
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   437
  let val (sign,tye,_) = foldr add_types (ctpairs0, (#sign(rep_thm th),[],0))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   438
      val tsig = #tsig(Sign.rep_sg sign);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   439
      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
   440
                         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
   441
      fun ctyp2 (ix,T) = (ix, ctyp_of sign T)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   442
  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
   443
  handle TERM _ =>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   444
           raise THM("cterm_instantiate: incompatible signatures",0,[th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   445
       | TYPE _ => raise THM("cterm_instantiate: types", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   446
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   447
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   448
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   449
(** theorem equality test is exported and used by BEST_FIRST **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   450
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   451
(*equality of theorems uses equality of signatures and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   452
  the a-convertible test for terms*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   453
fun eq_thm (th1,th2) =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   454
    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
   455
        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
   456
    in  Sign.eq_sg (sg1,sg2) andalso
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   457
        eq_set (shyps1, shyps2) andalso
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   458
        aconvs(hyps1,hyps2) andalso
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   459
        prop1 aconv prop2
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   460
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   461
1241
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   462
(*equality of theorems using similarity of signatures,
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   463
  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
   464
  version of this theory*)
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   465
fun same_thm (th1,th2) =
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   466
    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
   467
        and {sign=sg2, shyps=shyps2, hyps=hyps2, prop=prop2, ...} = rep_thm th2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   468
    in  Sign.same_sg (sg1,sg2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   469
        eq_set (shyps1, shyps2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   470
        aconvs(hyps1,hyps2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   471
        prop1 aconv prop2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   472
    end;
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   473
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
(*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
   475
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
   476
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   477
(*Useful "distance" function for BEST_FIRST*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   478
val size_of_thm = size_of_term o #prop o rep_thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   479
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   480
1194
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   481
(** 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
   482
    (some) type variable renaming **)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   483
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   484
 (* 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
   485
    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
   486
    in the term. *)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   487
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
   488
  | 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
   489
  | 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
   490
  | term_vars' _ = [];
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   491
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   492
fun forall_intr_vars th =
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   493
  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
   494
      val vars = distinct (term_vars' prop);
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   495
  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
   496
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   497
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
   498
    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
   499
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   500
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   501
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   502
(*** Meta-Rewriting Rules ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   503
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   504
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   505
val reflexive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   506
  let val cx = cterm_of Sign.proto_pure (Var(("x",0),TVar(("'a",0),logicS)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   507
  in Thm.reflexive cx end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   508
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   509
val symmetric_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   510
  let val xy = read_cterm Sign.proto_pure ("x::'a::logic == y",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   511
  in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   512
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   513
val transitive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   514
  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
   515
      val yz = read_cterm Sign.proto_pure ("y::'a::logic == z",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   516
      val xythm = Thm.assume xy and yzthm = Thm.assume yz
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   517
  in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   518
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   519
(** 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
   520
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   521
val refl_cimplies = reflexive (cterm_of Sign.proto_pure implies);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   522
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   523
(*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
   524
(*Do not rewrite flex-flex pairs*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   525
fun goals_conv pred cv =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   526
  let fun gconv i ct =
1703
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
   527
        let val (A,B) = dest_cimplies ct
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   528
            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
   529
                  Const("=?=",_)$_$_ => (reflexive A, i)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   530
                | _ => (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
   531
        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
   532
        handle TERM _ => reflexive ct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   533
  in gconv 1 end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   534
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   535
(*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
   536
fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   537
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   538
(*rewriting conversion*)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   539
fun rew_conv mode prover mss = rewrite_cterm mode mss prover;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   540
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   541
(*Rewrite a theorem*)
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   542
fun rewrite_rule []   th = th
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   543
  | rewrite_rule thms th =
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   544
	fconv_rule (rew_conv (true,false) (K(K None)) (Thm.mss_of thms)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   545
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   546
(*Rewrite the subgoals of a proof state (represented by a theorem) *)
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   547
fun rewrite_goals_rule []   th = th
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   548
  | rewrite_goals_rule thms th =
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   549
	fconv_rule (goals_conv (K true) 
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   550
		    (rew_conv (true,false) (K(K None))
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   551
		     (Thm.mss_of thms))) 
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
   552
	           th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   553
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   554
(*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
   555
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
   556
  if 0 < i  andalso  i <= nprems_of thm
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   557
  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
   558
  else raise THM("rewrite_goal_rule",i,[thm]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   559
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   560
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   561
(** Derived rules mainly for METAHYPS **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   562
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   563
(*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   564
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
   565
  let val {sign=signa, t=a, ...} = rep_cterm ca
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   566
      and combth = combination eqth (reflexive ca)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   567
      val {sign,prop,...} = rep_thm eqth
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   568
      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
   569
      val cterm = cterm_of (Sign.merge (sign,signa))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   570
  in  transitive (symmetric (beta_conversion (cterm (abst$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   571
           (transitive combth (beta_conversion (cterm (absu$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   572
  end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   573
  handle THM _ => raise THM("equal_abs_elim", 0, [eqth]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   574
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   575
(*Calling equal_abs_elim with multiple terms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   576
fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   577
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   578
local
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   579
  open Logic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   580
  val alpha = TVar(("'a",0), [])     (*  type ?'a::{}  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   581
  fun err th = raise THM("flexpair_inst: ", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   582
  fun flexpair_inst def th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   583
    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
   584
        val cterm = cterm_of sign
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   585
        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
   586
        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
   587
                   def
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   588
    in  equal_elim def' th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   589
    end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   590
    handle THM _ => err th | bind => err th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   591
in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   592
val flexpair_intr = flexpair_inst (symmetric flexpair_def)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   593
and flexpair_elim = flexpair_inst flexpair_def
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   594
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   595
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   596
(*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
   597
fun flexpair_abs_elim_list cts =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   598
    flexpair_intr o equal_abs_elim_list cts o flexpair_elim;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   599
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   600
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   601
(*** Some useful meta-theorems ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   602
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   603
(*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
   604
val asm_rl = trivial(read_cterm Sign.proto_pure ("PROP ?psi",propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   605
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   606
(*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
   607
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
   608
        ("PROP ?psi ==> PROP ?theta", propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   609
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   610
(*Generalized elim rule for one conclusion; cut_rl with reversed premises:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   611
     [| PROP V;  PROP V ==> PROP W |] ==> PROP W *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   612
val revcut_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   613
  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
   614
      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
   615
  in  standard (implies_intr V
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   616
                (implies_intr VW
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   617
                 (implies_elim (assume VW) (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   618
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   619
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   620
(*for deleting an unwanted assumption*)
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   621
val thin_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   622
  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
   623
      and W = read_cterm Sign.proto_pure ("PROP W", propT);
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   624
  in  standard (implies_intr V (implies_intr W (assume W)))
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   625
  end;
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   626
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   627
(* (!!x. PROP ?V) == PROP ?V       Allows removal of redundant parameters*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   628
val triv_forall_equality =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   629
  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
   630
      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
   631
      and x  = read_cterm Sign.proto_pure ("x", TFree("'a",logicS));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   632
  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
   633
                           (implies_intr V  (forall_intr x (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   634
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   635
1756
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   636
(* (PROP ?PhiA ==> PROP ?PhiB ==> PROP ?Psi) ==>
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   637
   (PROP ?PhiB ==> PROP ?PhiA ==> PROP ?Psi)
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   638
   `thm COMP swap_prems_rl' swaps the first two premises of `thm'
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   639
*)
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   640
val swap_prems_rl =
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   641
  let val cmajor = read_cterm Sign.proto_pure
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   642
            ("PROP PhiA ==> PROP PhiB ==> PROP Psi", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   643
      val major = assume cmajor;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   644
      val cminor1 = read_cterm Sign.proto_pure  ("PROP PhiA", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   645
      val minor1 = assume cminor1;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   646
      val cminor2 = read_cterm Sign.proto_pure  ("PROP PhiB", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   647
      val minor2 = assume cminor2;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   648
  in standard
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   649
       (implies_intr cmajor (implies_intr cminor2 (implies_intr cminor1
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   650
         (implies_elim (implies_elim major minor1) minor2))))
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   651
  end;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   652
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   653
end;
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   654
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
   655
open Drule;