src/Pure/drule.ML
author wenzelm
Wed, 15 Oct 1997 15:12:59 +0200
changeset 3872 a5839ecee7b8
parent 3766 8e1794c4e81b
child 3991 4cb2f2422695
permissions -rw-r--r--
tuned; prepare ext;
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
3766
8e1794c4e81b moved theory stuff (add_defs etc.) to theory.ML;
wenzelm
parents: 3653
diff changeset
     6
Derived rules and other operations on theorems.
0
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 =
3766
8e1794c4e81b moved theory stuff (add_defs etc.) to theory.ML;
wenzelm
parents: 3653
diff changeset
    12
sig
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    13
  val asm_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    14
  val assume_ax		: theory -> string -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    15
  val COMP		: thm * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    16
  val compose		: thm * int * thm -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    17
  val cprems_of		: thm -> cterm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    18
  val cterm_instantiate	: (cterm*cterm)list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    19
  val cut_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    20
  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
    21
  val equal_abs_elim_list: cterm list -> thm -> thm
3653
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
    22
  val equal_intr_rule   : thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    23
  val eq_thm		: thm * thm -> bool
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    24
  val same_thm		: thm * thm -> bool
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    25
  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
    26
  val flexpair_abs_elim_list: cterm list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    27
  val forall_intr_list	: cterm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    28
  val forall_intr_frees	: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    29
  val forall_intr_vars	: thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    30
  val forall_elim_list	: cterm list -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    31
  val forall_elim_var	: int -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    32
  val forall_elim_vars	: int -> thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    33
  val implies_elim_list	: thm -> thm list -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    34
  val implies_intr_list	: cterm list -> thm -> thm
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    35
  val dest_implies      : cterm -> cterm * cterm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    36
  val MRL		: thm list list * thm list -> thm list
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    37
  val MRS		: thm list * thm -> thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    38
  val read_instantiate	: (string*string)list -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
  val read_instantiate_sg: Sign.sg -> (string*string)list -> thm -> thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    40
  val read_insts	:
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    41
          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
    42
                  -> (indexname -> typ option) * (indexname -> sort option)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
    43
                  -> string list -> (string*string)list
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
    44
                  -> (indexname*ctyp)list * (cterm*cterm)list
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    45
  val reflexive_thm	: thm
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    46
  val refl_implies      : thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    47
  val revcut_rl		: thm
3555
5a720f6b9f38 added rewrite_thm;
wenzelm
parents: 3530
diff changeset
    48
  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
    49
        -> meta_simpset -> int -> thm -> thm
3575
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
    50
  val rewrite_goals_rule_aux: (meta_simpset -> thm -> thm option) -> thm list -> thm -> thm
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
    51
  val rewrite_rule_aux	: (meta_simpset -> thm -> thm option) -> thm list -> thm -> thm
3555
5a720f6b9f38 added rewrite_thm;
wenzelm
parents: 3530
diff changeset
    52
  val rewrite_thm	: bool * bool -> (meta_simpset -> thm -> thm option)
5a720f6b9f38 added rewrite_thm;
wenzelm
parents: 3530
diff changeset
    53
	-> meta_simpset -> thm -> thm
1460
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
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    59
  val skip_flexpairs	: cterm -> cterm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    60
  val standard		: thm -> thm
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    61
  val strip_imp_prems	: cterm -> cterm list
1756
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
    62
  val swap_prems_rl     : thm
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    63
  val symmetric_thm	: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    64
  val thin_rl		: thm
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    65
  val transitive_thm	: thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
  val triv_forall_equality: thm
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
  val types_sorts: thm -> (indexname-> typ option) * (indexname-> sort option)
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    68
  val zero_var_indexes	: thm -> thm
3766
8e1794c4e81b moved theory stuff (add_defs etc.) to theory.ML;
wenzelm
parents: 3653
diff changeset
    69
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
    71
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
    72
structure Drule : DRULE =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    75
(** 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
    76
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    77
(** SAME NAMES as in structure Logic: use compound identifiers! **)
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    78
1703
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
    79
(*dest_implies for cterms. Note T=prop below*)
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    80
fun dest_implies ct =
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    81
    case term_of ct of 
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    82
	(Const("==>", _) $ _ $ _) => 
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    83
	    let val (ct1,ct2) = dest_comb ct
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    84
	    in  (#2 (dest_comb ct1), ct2)  end	     
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    85
      | _ => raise TERM ("dest_implies", [term_of ct]) ;
1703
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
    86
e22ad43bab5f moved dest_cimplies to drule.ML; added adjust_maxidx
clasohm
parents: 1596
diff changeset
    87
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    88
(*Discard flexflex pairs; return a cterm*)
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    89
fun skip_flexpairs ct =
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    90
    case term_of ct of
1460
5a6f2aabd538 inserted tabs again
clasohm
parents: 1458
diff changeset
    91
	(Const("==>", _) $ (Const("=?=",_)$_$_) $ _) =>
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    92
	    skip_flexpairs (#2 (dest_implies ct))
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    93
      | _ => ct;
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    94
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    95
(* A1==>...An==>B  goes to  [A1,...,An], where B is not an implication *)
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    96
fun strip_imp_prems ct =
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    97
    let val (cA,cB) = dest_implies ct
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
    98
    in  cA :: strip_imp_prems cB  end
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
    99
    handle TERM _ => [];
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   100
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   101
(* A1==>...An==>B  goes to B, where B is not an implication *)
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   102
fun strip_imp_concl ct =
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   103
    case term_of ct of (Const("==>", _) $ _ $ _) => 
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   104
	strip_imp_concl (#2 (dest_comb ct))
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   105
  | _ => ct;
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   106
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   107
(*The premises of a theorem, as a cterm list*)
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   108
val cprems_of = strip_imp_prems o skip_flexpairs o cprop_of;
708
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   109
8422e50adce0 Pure/drule/cprems_of, cskip_flexpairs, cstrip_imp_prems: new cterm operations
lcp
parents: 668
diff changeset
   110
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   111
(** reading of instantiations **)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   112
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   113
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
   114
        | _ => 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
   115
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   116
fun absent ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   117
  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
   118
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   119
fun inst_failure ixn =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   120
  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
   121
952
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   122
(* 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
   123
   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
   124
   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
   125
   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
   126
*)
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   127
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
   128
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
   129
    fun split([],tvs,vs) = (tvs,vs)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   130
      | 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
   131
                  "'"::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
   132
                | 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
   133
    val (tvs,vs) = split(insts,[],[]);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   134
    fun readT((a,i),st) =
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   135
        let val ixn = ("'" ^ a,i);
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   136
            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
   137
            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
   138
        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
   139
           else inst_failure ixn
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   140
        end
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   141
    val tye = map readT tvs;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   142
    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
   143
        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
   144
                      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
   145
                    | None => absent ixn;
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   146
            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
   147
            val cts' = (ixn,T,ct)::cts
9e10962866b0 Removed an old bug which made some simultaneous instantiations fail if they
nipkow
parents: 949
diff changeset
   148
            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
   149
            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
   150
        in (map inst cts',tye2 @ tye,used') end
949
83c588d6fee9 Changed treatment of during type inference internally generated type
nipkow
parents: 922
diff changeset
   151
    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
   152
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
   153
    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
   154
end;
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   155
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   156
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   157
(*** Find the type (sort) associated with a (T)Var or (T)Free in a term
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
     Used for establishing default types (of variables) and sorts (of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   159
     type variables) when reading another term.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   160
     Index -1 indicates that a (T)Free rather than a (T)Var is wanted.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
fun types_sorts thm =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
    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
   165
        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
   166
        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
   167
        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
   168
        val tvars = term_tvars big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   169
        val tfrees = term_tfrees big;
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   170
        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
   171
        fun sort(a,i) = if i<0 then assoc(tfrees,a) else assoc(tvars,(a,i));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
    in (typ,sort) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
(** Standardization of rules **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
(*Generalization over a list of variables, IGNORING bad ones*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   177
fun forall_intr_list [] th = th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
  | forall_intr_list (y::ys) th =
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   179
        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
   180
        in  forall_intr y gth   handle THM _ =>  gth  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   182
(*Generalization over all suitable Free variables*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
fun forall_intr_frees th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   185
    in  forall_intr_list
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   186
         (map (cterm_of sign) (sort atless (term_frees prop)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
         th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
(*Replace outermost quantified variable by Var of given index.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
    Could clash with Vars already present.*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   192
fun forall_elim_var i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
    let val {prop,sign,...} = rep_thm th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
    in case prop of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   195
          Const("all",_) $ Abs(a,T,_) =>
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   196
              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
   197
        | _ => raise THM("forall_elim_var", i, [th])
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   200
(*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
   201
fun forall_elim_vars i th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   202
    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
   203
        handle THM _ => th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   204
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
(*Specialization over a list of cterms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
fun forall_elim_list cts th = foldr (uncurry forall_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
(* maps [A1,...,An], B   to   [| A1;...;An |] ==> B  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
fun implies_intr_list cAs th = foldr (uncurry implies_intr) (cAs,th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   211
(* maps [| A1;...;An |] ==> B and [A1,...,An]   to   B *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   212
fun implies_elim_list impth ths = foldl (uncurry implies_elim) (impth,ths);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   213
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   214
(*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
   215
fun zero_var_indexes th =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
    let val {prop,sign,...} = rep_thm th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   217
        val vars = term_vars prop
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
        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
   219
        val inrs = add_term_tvars(prop,[]);
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   220
        val nms' = rev(foldl add_new_id ([], map (#1 o #1) inrs));
2266
82aef6857c5b Replaced map...~~ by ListPair.map
paulson
parents: 2180
diff changeset
   221
        val tye = ListPair.map (fn ((v,rs),a) => (v, TVar((a,0),rs)))
82aef6857c5b Replaced map...~~ by ListPair.map
paulson
parents: 2180
diff changeset
   222
	             (inrs, nms')
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   223
        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
   224
        fun varpairs([],[]) = []
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   225
          | 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
   226
                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
   227
                in (cterm_of sign (Var(v,T')),
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   228
                    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
   229
                end
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   230
          | varpairs _ = raise TERM("varpairs", []);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   231
    in instantiate (ctye, varpairs(vars,rev bs)) th end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   232
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   233
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   234
(*Standard form of object-rule: no hypotheses, Frees, or outer quantifiers;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
    all generality expressed by Vars having index 0.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   236
fun standard th =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   237
  let val {maxidx,...} = rep_thm th
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   238
  in
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   239
    th |> implies_intr_hyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   240
       |> forall_intr_frees |> forall_elim_vars (maxidx + 1)
1439
1f5949a43e82 improved printing of errors in 'defs';
wenzelm
parents: 1435
diff changeset
   241
       |> Thm.strip_shyps |> Thm.implies_intr_shyps
1412
2ab32768c996 Now "standard" compresses theorems (for sharing).
paulson
parents: 1241
diff changeset
   242
       |> zero_var_indexes |> Thm.varifyT |> Thm.compress
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   243
  end;
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   244
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   245
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   246
(*Assume a new formula, read following the same conventions as axioms.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   247
  Generalizes over Free variables,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
  creates the assumption, and then strips quantifiers.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   249
  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
   250
             [ !(A,P,a)[| ALL x:A. P(x) |] ==> [| P(a) |] ]    *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   251
fun assume_ax thy sP =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
    let val sign = sign_of thy
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   253
        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
   254
                         (sP, propT)))
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   255
    in forall_elim_vars 0 (assume (cterm_of sign prop))  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   256
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   257
(*Resolution: exactly one resolvent must be produced.*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   258
fun tha RSN (i,thb) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   259
  case Sequence.chop (2, biresolution false [(false,tha)] i thb) of
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   260
      ([th],_) => th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   261
    | ([],_)   => raise THM("RSN: no unifiers", i, [tha,thb])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   262
    |      _   => raise THM("RSN: multiple unifiers", i, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   263
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   264
(*resolution: P==>Q, Q==>R gives P==>R. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   265
fun tha RS thb = tha RSN (1,thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   266
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   267
(*For joining lists of rules*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   268
fun thas RLN (i,thbs) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
  let val resolve = biresolution false (map (pair false) thas) i
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
      fun resb thb = Sequence.list_of_s (resolve thb) handle THM _ => []
2672
85d7e800d754 Replaced "flat" by the Basis Library function List.concat
paulson
parents: 2266
diff changeset
   271
  in  List.concat (map resb thbs)  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   272
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   273
fun thas RL thbs = thas RLN (1,thbs);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
11
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   275
(*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
   276
  makes proof trees*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   277
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
   278
  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
   279
        | 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
   280
  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
   281
d0e17c42dbb4 Added MRS, MRL from ZF/ROOT.ML. These support forward proof, resolving a
lcp
parents: 0
diff changeset
   282
(*As above, but for rule lists*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   283
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
   284
  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
   285
        | 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
   286
  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
   287
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   288
(*compose Q and [...,Qi,Q(i+1),...]==>R to [...,Q(i+1),...]==>R
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   289
  with no lifting or renaming!  Q may contain ==> or meta-quants
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
  ALWAYS deletes premise i *)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   291
fun compose(tha,i,thb) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
    Sequence.list_of_s (bicompose false (false,tha,0) i thb);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
(*compose Q and [Q1,Q2,...,Qk]==>R to [Q2,...,Qk]==>R getting unique result*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
fun tha COMP thb =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   296
    case compose(tha,1,thb) of
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   297
        [th] => th
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
      | _ =>   raise THM("COMP", 1, [tha,thb]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   299
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   300
(*Instantiate theorem th, reading instantiations under signature sg*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
fun read_instantiate_sg sg sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
    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
   303
        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
   304
    in  instantiate (read_insts sg ts ts used sinsts) th  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   305
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   306
(*Instantiate theorem th, reading instantiations under theory of th*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
fun read_instantiate sinsts th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
    read_instantiate_sg (#sign (rep_thm th)) sinsts th;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   311
(*Left-to-right replacements: tpairs = [...,(vi,ti),...].
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
  Instantiates distinct Vars by terms, inferring type instantiations. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
local
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   314
  fun add_types ((ct,cu), (sign,tye,maxidx)) =
2152
76d5ed939545 Now uses Int.max instead of max
paulson
parents: 2004
diff changeset
   315
    let val {sign=signt, t=t, T= T, maxidx=maxt,...} = rep_cterm ct
76d5ed939545 Now uses Int.max instead of max
paulson
parents: 2004
diff changeset
   316
        and {sign=signu, t=u, T= U, maxidx=maxu,...} = rep_cterm cu;
76d5ed939545 Now uses Int.max instead of max
paulson
parents: 2004
diff changeset
   317
        val maxi = Int.max(maxidx, Int.max(maxt, maxu));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
        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
   319
        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
   320
          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
   321
    in  (sign', tye', maxi')  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   322
in
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   323
fun cterm_instantiate ctpairs0 th =
1435
aefcd255ed4a Removed bug in type unification. Negative indexes are not used any longer.
nipkow
parents: 1412
diff changeset
   324
  let val (sign,tye,_) = foldr add_types (ctpairs0, (#sign(rep_thm th),[],0))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   325
      val tsig = #tsig(Sign.rep_sg sign);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
      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
   327
                         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
   328
      fun ctyp2 (ix,T) = (ix, ctyp_of sign T)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   329
  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
   330
  handle TERM _ =>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   331
           raise THM("cterm_instantiate: incompatible signatures",0,[th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   332
       | TYPE _ => raise THM("cterm_instantiate: types", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   333
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   334
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   335
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   336
(** theorem equality test is exported and used by BEST_FIRST **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   337
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   338
(*equality of theorems uses equality of signatures and
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   339
  the a-convertible test for terms*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   340
fun eq_thm (th1,th2) =
1218
59ed8ef1a3a1 modified pretty_thm, standard, eq_thm to handle shyps;
wenzelm
parents: 1194
diff changeset
   341
    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
   342
        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
   343
    in  Sign.eq_sg (sg1,sg2) andalso
2180
934572a94139 Removal of polymorphic equality via mem, subset, eq_set, etc
paulson
parents: 2152
diff changeset
   344
        eq_set_sort (shyps1, shyps2) andalso
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   345
        aconvs(hyps1,hyps2) andalso
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   346
        prop1 aconv prop2
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   347
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   348
1241
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   349
(*equality of theorems using similarity of signatures,
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   350
  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
   351
  version of this theory*)
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   352
fun same_thm (th1,th2) =
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   353
    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
   354
        and {sign=sg2, shyps=shyps2, hyps=hyps2, prop=prop2, ...} = rep_thm th2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   355
    in  Sign.same_sg (sg1,sg2) andalso
2180
934572a94139 Removal of polymorphic equality via mem, subset, eq_set, etc
paulson
parents: 2152
diff changeset
   356
        eq_set_sort (shyps1, shyps2) andalso
1241
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   357
        aconvs(hyps1,hyps2) andalso
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   358
        prop1 aconv prop2
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   359
    end;
bfc93c86f0a1 added same_sg and same_thm
clasohm
parents: 1237
diff changeset
   360
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   361
(*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
   362
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
   363
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
(*Useful "distance" function for BEST_FIRST*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
val size_of_thm = size_of_term o #prop o rep_thm;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   367
1194
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   368
(** 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
   369
    (some) type variable renaming **)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   370
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   371
 (* 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
   372
    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
   373
    in the term. *)
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   374
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
   375
  | 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
   376
  | 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
   377
  | term_vars' _ = [];
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   378
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   379
fun forall_intr_vars th =
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   380
  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
   381
      val vars = distinct (term_vars' prop);
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   382
  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
   383
1237
45ac644b0052 adapted to new version of shyps-stuff;
wenzelm
parents: 1218
diff changeset
   384
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
   385
    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
   386
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   387
563ecd14c1d8 Added weak_eq_thm and forall_intr_vars (thanks to Mark Staples)
lcp
parents: 952
diff changeset
   388
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   389
(*** Meta-Rewriting Rules ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   391
val reflexive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   392
  let val cx = cterm_of Sign.proto_pure (Var(("x",0),TVar(("'a",0),logicS)))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   393
  in Thm.reflexive cx end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   394
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   395
val symmetric_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   396
  let val xy = read_cterm Sign.proto_pure ("x::'a::logic == y",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   397
  in standard(Thm.implies_intr_hyps(Thm.symmetric(Thm.assume xy))) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   398
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   399
val transitive_thm =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   400
  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
   401
      val yz = read_cterm Sign.proto_pure ("y::'a::logic == z",propT)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   402
      val xythm = Thm.assume xy and yzthm = Thm.assume yz
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   403
  in standard(Thm.implies_intr yz (Thm.transitive xythm yzthm)) end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   404
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   405
(** 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
   406
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   407
val refl_implies = reflexive (cterm_of Sign.proto_pure implies);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   409
(*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
   410
(*Do not rewrite flex-flex pairs*)
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   411
fun goals_conv pred cv =
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   412
  let fun gconv i ct =
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   413
        let val (A,B) = dest_implies ct
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   414
            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
   415
                  Const("=?=",_)$_$_ => (reflexive A, i)
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   416
                | _ => (if pred i then cv A else reflexive A, i+1)
2004
3411fe560611 New operations on cterms. Now same names as in Logic
paulson
parents: 1906
diff changeset
   417
        in  combination (combination refl_implies thA) (gconv j B) end
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   418
        handle TERM _ => reflexive ct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   419
  in gconv 1 end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   420
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
(*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
   422
fun fconv_rule cv th = equal_elim (cv (cprop_of th)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   424
(*rewriting conversion*)
229
4002c4cd450c Pure: MAJOR CHANGE. Moved ML types ctyp and cterm and their associated
lcp
parents: 214
diff changeset
   425
fun rew_conv mode prover mss = rewrite_cterm mode mss prover;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   426
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   427
(*Rewrite a theorem*)
3575
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   428
fun rewrite_rule_aux _ []   th = th
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   429
  | rewrite_rule_aux prover thms th =
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   430
      fconv_rule (rew_conv (true,false) prover (Thm.mss_of thms)) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   431
3555
5a720f6b9f38 added rewrite_thm;
wenzelm
parents: 3530
diff changeset
   432
fun rewrite_thm mode prover mss = fconv_rule (rew_conv mode prover mss);
5a720f6b9f38 added rewrite_thm;
wenzelm
parents: 3530
diff changeset
   433
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   434
(*Rewrite the subgoals of a proof state (represented by a theorem) *)
3575
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   435
fun rewrite_goals_rule_aux _ []   th = th
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   436
  | rewrite_goals_rule_aux prover thms th =
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   437
      fconv_rule (goals_conv (K true) (rew_conv (true, true) prover
4e9beacb5339 improved rewrite_thm / rewrite_goals to handle conditional eqns;
wenzelm
parents: 3555
diff changeset
   438
        (Thm.mss_of thms))) th;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   439
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   440
(*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
   441
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
   442
  if 0 < i  andalso  i <= nprems_of thm
ed6a3e2b1a33 added new parameter to the simplification tactics which indicates if
nipkow
parents: 211
diff changeset
   443
  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
   444
  else raise THM("rewrite_goal_rule",i,[thm]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   445
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   446
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   447
(** Derived rules mainly for METAHYPS **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   448
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   449
(*Given the term "a", takes (%x.t)==(%x.u) to t[a/x]==u[a/x]*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   450
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
   451
  let val {sign=signa, t=a, ...} = rep_cterm ca
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   452
      and combth = combination eqth (reflexive ca)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   453
      val {sign,prop,...} = rep_thm eqth
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   454
      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
   455
      val cterm = cterm_of (Sign.merge (sign,signa))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   456
  in  transitive (symmetric (beta_conversion (cterm (abst$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   457
           (transitive combth (beta_conversion (cterm (absu$a))))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   458
  end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   459
  handle THM _ => raise THM("equal_abs_elim", 0, [eqth]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   460
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   461
(*Calling equal_abs_elim with multiple terms*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   462
fun equal_abs_elim_list cts th = foldr (uncurry equal_abs_elim) (rev cts, th);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   463
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   464
local
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   465
  val alpha = TVar(("'a",0), [])     (*  type ?'a::{}  *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   466
  fun err th = raise THM("flexpair_inst: ", 0, [th])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   467
  fun flexpair_inst def th =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   468
    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
   469
        val cterm = cterm_of sign
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   470
        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
   471
        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
   472
                   def
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   473
    in  equal_elim def' th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
    end
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   475
    handle THM _ => err th | bind => err th
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   476
in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   477
val flexpair_intr = flexpair_inst (symmetric flexpair_def)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   478
and flexpair_elim = flexpair_inst flexpair_def
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   479
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   480
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   481
(*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
   482
fun flexpair_abs_elim_list cts =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   483
    flexpair_intr o equal_abs_elim_list cts o flexpair_elim;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   484
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   485
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   486
(*** Some useful meta-theorems ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   487
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   488
(*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
   489
val asm_rl = trivial(read_cterm Sign.proto_pure ("PROP ?psi",propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   490
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   491
(*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
   492
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
   493
        ("PROP ?psi ==> PROP ?theta", propT));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   494
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   495
(*Generalized elim rule for one conclusion; cut_rl with reversed premises:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   496
     [| PROP V;  PROP V ==> PROP W |] ==> PROP W *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   497
val revcut_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   498
  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
   499
      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
   500
  in  standard (implies_intr V
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   501
                (implies_intr VW
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   502
                 (implies_elim (assume VW) (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   503
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   504
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   505
(*for deleting an unwanted assumption*)
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   506
val thin_rl =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   507
  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
   508
      and W = read_cterm Sign.proto_pure ("PROP W", propT);
668
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   509
  in  standard (implies_intr V (implies_intr W (assume W)))
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   510
  end;
0d0923eb0f0d Pure/drule/thin_rl: new
lcp
parents: 655
diff changeset
   511
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   512
(* (!!x. PROP ?V) == PROP ?V       Allows removal of redundant parameters*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   513
val triv_forall_equality =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 843
diff changeset
   514
  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
   515
      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
   516
      and x  = read_cterm Sign.proto_pure ("x", TFree("'a",logicS));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   517
  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
   518
                           (implies_intr V  (forall_intr x (assume V))))
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   519
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   520
1756
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   521
(* (PROP ?PhiA ==> PROP ?PhiB ==> PROP ?Psi) ==>
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   522
   (PROP ?PhiB ==> PROP ?PhiA ==> PROP ?Psi)
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   523
   `thm COMP swap_prems_rl' swaps the first two premises of `thm'
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   524
*)
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   525
val swap_prems_rl =
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   526
  let val cmajor = read_cterm Sign.proto_pure
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   527
            ("PROP PhiA ==> PROP PhiB ==> PROP Psi", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   528
      val major = assume cmajor;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   529
      val cminor1 = read_cterm Sign.proto_pure  ("PROP PhiA", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   530
      val minor1 = assume cminor1;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   531
      val cminor2 = read_cterm Sign.proto_pure  ("PROP PhiB", propT);
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   532
      val minor2 = assume cminor2;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   533
  in standard
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   534
       (implies_intr cmajor (implies_intr cminor2 (implies_intr cminor1
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   535
         (implies_elim (implies_elim major minor1) minor2))))
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   536
  end;
978ee7ededdd Added swap_prems_rl
nipkow
parents: 1703
diff changeset
   537
3653
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   538
(* [| PROP ?phi ==> PROP ?psi; PROP ?psi ==> PROP ?phi |]
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   539
   ==> PROP ?phi == PROP ?psi
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   540
   Introduction rule for == using ==> not meta-hyps.
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   541
*)
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   542
val equal_intr_rule =
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   543
  let val PQ = read_cterm Sign.proto_pure ("PROP phi ==> PROP psi", propT)
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   544
      and QP = read_cterm Sign.proto_pure ("PROP psi ==> PROP phi", propT)
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   545
  in  equal_intr (assume PQ) (assume QP)
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   546
      |> implies_intr QP
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   547
      |> implies_intr PQ
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   548
      |> standard
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   549
  end;
6d5b3d5ff132 Added Larry's equal_intr_rule
nipkow
parents: 3575
diff changeset
   550
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   551
end;
252
7532f95d7f44 removed eq_sg, pprint_sg, print_sg (now in sign.ML);
wenzelm
parents: 229
diff changeset
   552
1499
01fdd1ea6324 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
   553
open Drule;