src/Pure/logic.ML
author wenzelm
Tue, 21 Mar 2006 12:18:13 +0100
changeset 19304 15f001295a7b
parent 19125 59b26248547b
child 19391 4812d28c90a6
permissions -rw-r--r--
remove (op =); tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
     1
(*  Title:      Pure/logic.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   Cambridge University 1992
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
     6
Abstract syntax operations of the Pure meta-logic.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
     9
signature LOGIC =
4345
7e9436ffb813 tuned term order;
wenzelm
parents: 4318
diff changeset
    10
sig
18181
wenzelm
parents: 18029
diff changeset
    11
  val is_all: term -> bool
18762
9098c92a945f added dest_all;
wenzelm
parents: 18499
diff changeset
    12
  val dest_all: term -> typ * term
18181
wenzelm
parents: 18029
diff changeset
    13
  val mk_equals: term * term -> term
wenzelm
parents: 18029
diff changeset
    14
  val dest_equals: term -> term * term
wenzelm
parents: 18029
diff changeset
    15
  val is_equals: term -> bool
wenzelm
parents: 18029
diff changeset
    16
  val mk_implies: term * term -> term
wenzelm
parents: 18029
diff changeset
    17
  val dest_implies: term -> term * term
wenzelm
parents: 18029
diff changeset
    18
  val is_implies: term -> bool
wenzelm
parents: 18029
diff changeset
    19
  val list_implies: term list * term -> term
wenzelm
parents: 18029
diff changeset
    20
  val strip_imp_prems: term -> term list
wenzelm
parents: 18029
diff changeset
    21
  val strip_imp_concl: term -> term
wenzelm
parents: 18029
diff changeset
    22
  val strip_prems: int * term list * term -> term list * term
wenzelm
parents: 18029
diff changeset
    23
  val count_prems: term * int -> int
wenzelm
parents: 18029
diff changeset
    24
  val nth_prem: int * term -> term
19125
59b26248547b simplified Pure conjunction, based on actual const;
wenzelm
parents: 19103
diff changeset
    25
  val conjunction: term
18181
wenzelm
parents: 18029
diff changeset
    26
  val mk_conjunction: term * term -> term
12757
b76a4376cfcb added mk_conjunction_list;
wenzelm
parents: 12137
diff changeset
    27
  val mk_conjunction_list: term list -> term
18499
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
    28
  val mk_conjunction_list2: term list list -> term
18469
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
    29
  val dest_conjunction: term -> term * term
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
    30
  val dest_conjunctions: term -> term list
18181
wenzelm
parents: 18029
diff changeset
    31
  val strip_horn: term -> term list * term
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    32
  val dest_type: term -> typ
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    33
  val const_of_class: class -> string
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    34
  val class_of_const: string -> class
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    35
  val mk_inclass: typ * class -> term
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    36
  val dest_inclass: term -> typ * class
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    37
  val dest_def: Pretty.pp -> (term -> bool) -> (string -> bool) -> (string -> bool) ->
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    38
    term -> (term * term) * term
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    39
  val abs_def: term -> term * term
18181
wenzelm
parents: 18029
diff changeset
    40
  val mk_cond_defpair: term list -> term * term -> string * term
wenzelm
parents: 18029
diff changeset
    41
  val mk_defpair: term * term -> string * term
wenzelm
parents: 18029
diff changeset
    42
  val mk_type: typ -> term
wenzelm
parents: 18029
diff changeset
    43
  val protectC: term
wenzelm
parents: 18029
diff changeset
    44
  val protect: term -> term
wenzelm
parents: 18029
diff changeset
    45
  val unprotect: term -> term
wenzelm
parents: 18029
diff changeset
    46
  val occs: term * term -> bool
wenzelm
parents: 18029
diff changeset
    47
  val close_form: term -> term
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    48
  val combound: term * int * int -> term
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
    49
  val rlist_abs: (string * typ) list * term -> term
18181
wenzelm
parents: 18029
diff changeset
    50
  val incr_indexes: typ list * int -> term -> term
wenzelm
parents: 18029
diff changeset
    51
  val incr_tvar: int -> typ -> typ
wenzelm
parents: 18029
diff changeset
    52
  val lift_abs: int -> term -> term -> term
wenzelm
parents: 18029
diff changeset
    53
  val lift_all: int -> term -> term -> term
wenzelm
parents: 18029
diff changeset
    54
  val strip_assums_hyp: term -> term list
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
    55
  val strip_assums_concl: term -> term
18181
wenzelm
parents: 18029
diff changeset
    56
  val strip_params: term -> (string * typ) list
wenzelm
parents: 18029
diff changeset
    57
  val has_meta_prems: term -> int -> bool
wenzelm
parents: 18029
diff changeset
    58
  val flatten_params: int -> term -> term
wenzelm
parents: 18029
diff changeset
    59
  val auto_rename: bool ref
wenzelm
parents: 18029
diff changeset
    60
  val set_rename_prefix: string -> unit
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
  val list_rename_params: string list * term -> term
18181
wenzelm
parents: 18029
diff changeset
    62
  val assum_pairs: int * term -> (term*term)list
wenzelm
parents: 18029
diff changeset
    63
  val varify: term -> term
wenzelm
parents: 18029
diff changeset
    64
  val unvarify: term -> term
wenzelm
parents: 18029
diff changeset
    65
  val get_goal: term -> int -> term
wenzelm
parents: 18029
diff changeset
    66
  val goal_params: term -> int -> term * term list
wenzelm
parents: 18029
diff changeset
    67
  val prems_of_goal: term -> int -> term list
wenzelm
parents: 18029
diff changeset
    68
  val concl_of_goal: term -> int -> term
4345
7e9436ffb813 tuned term order;
wenzelm
parents: 4318
diff changeset
    69
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
1500
b2de3b3277b8 Elimination of fully-functorial style.
paulson
parents: 1460
diff changeset
    71
structure Logic : LOGIC =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
struct
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
    73
4345
7e9436ffb813 tuned term order;
wenzelm
parents: 4318
diff changeset
    74
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
(*** Abstract syntax operations on the meta-connectives ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
5041
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    77
(** all **)
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    78
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    79
fun is_all (Const ("all", _) $ _) = true
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    80
  | is_all _ = false;
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    81
18762
9098c92a945f added dest_all;
wenzelm
parents: 18499
diff changeset
    82
fun dest_all (Const ("all", Type ("fun", [Type ("fun", [T, _]), _])) $ A) = (T, A)
9098c92a945f added dest_all;
wenzelm
parents: 18499
diff changeset
    83
  | dest_all t = raise TERM ("dest_all", [t]);
9098c92a945f added dest_all;
wenzelm
parents: 18499
diff changeset
    84
9098c92a945f added dest_all;
wenzelm
parents: 18499
diff changeset
    85
5041
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
    86
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
(** equality **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
1835
07eee14f5bd4 Restored warning comment
paulson
parents: 1500
diff changeset
    89
(*Make an equality.  DOES NOT CHECK TYPE OF u*)
64
0bbe5d86cb38 logic/mk_equals,mk_flexpair: now calls fastype_of instead of type_of.
lcp
parents: 0
diff changeset
    90
fun mk_equals(t,u) = equals(fastype_of t) $ t $ u;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
fun dest_equals (Const("==",_) $ t $ u)  =  (t,u)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
  | dest_equals t = raise TERM("dest_equals", [t]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
637
b344bf624143 added is_equals: term -> bool;
wenzelm
parents: 585
diff changeset
    95
fun is_equals (Const ("==", _) $ _ $ _) = true
b344bf624143 added is_equals: term -> bool;
wenzelm
parents: 585
diff changeset
    96
  | is_equals _ = false;
b344bf624143 added is_equals: term -> bool;
wenzelm
parents: 585
diff changeset
    97
b344bf624143 added is_equals: term -> bool;
wenzelm
parents: 585
diff changeset
    98
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
(** implies **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
fun mk_implies(A,B) = implies $ A $ B;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
fun dest_implies (Const("==>",_) $ A $ B)  =  (A,B)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
  | dest_implies A = raise TERM("dest_implies", [A]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
5041
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
   106
fun is_implies (Const ("==>", _) $ _ $ _) = true
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
   107
  | is_implies _ = false;
a1d0a6d555cd Goals may now contain assumptions, which are not returned.
nipkow
parents: 4822
diff changeset
   108
4822
2733e21814fe added mk_cond_defpair, mk_defpair;
wenzelm
parents: 4693
diff changeset
   109
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
(** nested implications **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
(* [A1,...,An], B  goes to  A1==>...An==>B  *)
18181
wenzelm
parents: 18029
diff changeset
   113
fun list_implies ([], B) = B
wenzelm
parents: 18029
diff changeset
   114
  | list_implies (A::As, B) = implies $ A $ list_implies(As,B);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   116
(* A1==>...An==>B  goes to  [A1,...,An], where B is not an implication *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
fun strip_imp_prems (Const("==>", _) $ A $ B) = A :: strip_imp_prems B
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
  | strip_imp_prems _ = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
(* A1==>...An==>B  goes to B, where B is not an implication *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   121
fun strip_imp_concl (Const("==>", _) $ A $ B) = strip_imp_concl B
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
  | strip_imp_concl A = A : term;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
(*Strip and return premises: (i, [], A1==>...Ai==>B)
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   125
    goes to   ([Ai, A(i-1),...,A1] , B)         (REVERSED)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   126
  if  i<0 or else i too big then raises  TERM*)
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   127
fun strip_prems (0, As, B) = (As, B)
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   128
  | strip_prems (i, As, Const("==>", _) $ A $ B) =
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   129
        strip_prems (i-1, A::As, B)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
  | strip_prems (_, As, A) = raise TERM("strip_prems", A::As);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
16130
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   132
(*Count premises -- quicker than (length o strip_prems) *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
fun count_prems (Const("==>", _) $ A $ B, n) = count_prems (B,n+1)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
  | count_prems (_,n) = n;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
16130
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   136
(*Select Ai from A1 ==>...Ai==>B*)
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   137
fun nth_prem (1, Const ("==>", _) $ A $ _) = A
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   138
  | nth_prem (i, Const ("==>", _) $ _ $ B) = nth_prem (i - 1, B)
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   139
  | nth_prem (_, A) = raise TERM ("nth_prem", [A]);
38b111451155 added nth_prem;
wenzelm
parents: 15596
diff changeset
   140
13659
3cf622f6b0b2 Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents: 12902
diff changeset
   141
(*strip a proof state (Horn clause):
3cf622f6b0b2 Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents: 12902
diff changeset
   142
  B1 ==> ... Bn ==> C   goes to   ([B1, ..., Bn], C)    *)
3cf622f6b0b2 Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents: 12902
diff changeset
   143
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
3cf622f6b0b2 Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents: 12902
diff changeset
   144
4822
2733e21814fe added mk_cond_defpair, mk_defpair;
wenzelm
parents: 4693
diff changeset
   145
12137
6123958975b8 added mk_conjunction;
wenzelm
parents: 10816
diff changeset
   146
(** conjunction **)
6123958975b8 added mk_conjunction;
wenzelm
parents: 10816
diff changeset
   147
19125
59b26248547b simplified Pure conjunction, based on actual const;
wenzelm
parents: 19103
diff changeset
   148
val conjunction = Const ("ProtoPure.conjunction", propT --> propT --> propT);
59b26248547b simplified Pure conjunction, based on actual const;
wenzelm
parents: 19103
diff changeset
   149
18499
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   150
(*A && B*)
19125
59b26248547b simplified Pure conjunction, based on actual const;
wenzelm
parents: 19103
diff changeset
   151
fun mk_conjunction (A, B) = conjunction $ A $ B;
12137
6123958975b8 added mk_conjunction;
wenzelm
parents: 10816
diff changeset
   152
18499
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   153
(*A && B && C -- improper*)
12757
b76a4376cfcb added mk_conjunction_list;
wenzelm
parents: 12137
diff changeset
   154
fun mk_conjunction_list [] = Term.all propT $ Abs ("dummy", propT, mk_implies (Bound 0, Bound 0))
b76a4376cfcb added mk_conjunction_list;
wenzelm
parents: 12137
diff changeset
   155
  | mk_conjunction_list ts = foldr1 mk_conjunction ts;
12137
6123958975b8 added mk_conjunction;
wenzelm
parents: 10816
diff changeset
   156
18499
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   157
(*(A1 && B1 && C1) && (A2 && B2 && C2 && D2) && A3 && B3 -- improper*)
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   158
fun mk_conjunction_list2 tss =
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   159
  mk_conjunction_list (map mk_conjunction_list (filter_out null tss));
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   160
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   161
(*A && B*)
19125
59b26248547b simplified Pure conjunction, based on actual const;
wenzelm
parents: 19103
diff changeset
   162
fun dest_conjunction (Const ("ProtoPure.conjunction", _) $ A $ B) = (A, B)
18469
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   163
  | dest_conjunction t = raise TERM ("dest_conjunction", [t]);
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   164
18499
567370efb6d7 added mk_conjunction_list2;
wenzelm
parents: 18469
diff changeset
   165
(*((A && B) && C) && D && E -- flat*)
18469
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   166
fun dest_conjunctions t =
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   167
  (case try dest_conjunction t of
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   168
    NONE => [t]
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   169
  | SOME (A, B) => dest_conjunctions A @ dest_conjunctions B);
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   170
324245a561b5 mk_conjunction: proper treatment of bounds;
wenzelm
parents: 18248
diff changeset
   171
12137
6123958975b8 added mk_conjunction;
wenzelm
parents: 10816
diff changeset
   172
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   173
(** types as terms **)
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   174
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   175
fun mk_type ty = Const ("TYPE", itselfT ty);
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   176
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   177
fun dest_type (Const ("TYPE", Type ("itself", [ty]))) = ty
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   178
  | dest_type t = raise TERM ("dest_type", [t]);
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   179
4822
2733e21814fe added mk_cond_defpair, mk_defpair;
wenzelm
parents: 4693
diff changeset
   180
447
d1f827fa0a18 changed comment only;
wenzelm
parents: 398
diff changeset
   181
(** class constraints **)
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   182
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   183
val classN = "_class";
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   184
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   185
val const_of_class = suffix classN;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   186
fun class_of_const c = unsuffix classN c
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   187
  handle Fail _ => raise TERM ("class_of_const: bad name " ^ quote c, []);
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   188
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   189
fun mk_inclass (ty, c) =
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   190
  Const (const_of_class c, itselfT ty --> propT) $ mk_type ty;
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   191
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   192
fun dest_inclass (t as Const (c_class, _) $ ty) =
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   193
      ((dest_type ty, class_of_const c_class)
398
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   194
        handle TERM _ => raise TERM ("dest_inclass", [t]))
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   195
  | dest_inclass t = raise TERM ("dest_inclass", [t]);
41f279b477e2 added mk_type, dest_type, mk_inclass, dest_inclass (for axclasses);
wenzelm
parents: 210
diff changeset
   196
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   197
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   198
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   199
(** definitions **)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   200
19103
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   201
fun term_kind (Const _) = "existing constant "
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   202
  | term_kind (Free _) = "free variable "
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   203
  | term_kind (Bound _) = "bound variable "
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   204
  | term_kind _ = "";
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   205
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   206
(*c x == t[x] to !!x. c x == t[x]*)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   207
fun dest_def pp check_head is_fixed is_fixedT eq =
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   208
  let
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   209
    fun err msg = raise TERM (msg, [eq]);
19071
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   210
    val eq_vars = Term.strip_all_vars eq;
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   211
    val eq_body = Term.strip_all_body eq;
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   212
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   213
    val display_terms = commas_quote o map (Pretty.string_of_term pp o Syntax.bound_vars eq_vars);
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   214
    val display_types = commas_quote o map (Pretty.string_of_typ pp);
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   215
19071
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   216
    val (raw_lhs, rhs) = dest_equals eq_body handle TERM _ => err "Not a meta-equality (==)";
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   217
    val lhs = Envir.beta_eta_contract raw_lhs;
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   218
    val (head, args) = Term.strip_comb lhs;
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   219
    val head_tfrees = Term.add_tfrees head [];
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   220
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   221
    fun check_arg (Bound _) = true
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   222
      | check_arg (Free (x, _)) = not (is_fixed x)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   223
      | check_arg (Const ("TYPE", Type ("itself", [TFree _]))) = true
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   224
      | check_arg _ = false;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   225
    fun close_arg (Bound _) t = t
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   226
      | close_arg x t = Term.all (Term.fastype_of x) $ lambda x t;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   227
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   228
    val lhs_bads = filter_out check_arg args;
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18938
diff changeset
   229
    val lhs_dups = duplicates (op aconv) args;
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   230
    val rhs_extras = Term.fold_aterms (fn v as Free (x, _) =>
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   231
      if is_fixed x orelse member (op aconv) args v then I
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   232
      else insert (op aconv) v | _ => I) rhs [];
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   233
    val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) =>
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   234
      if is_fixedT a orelse member (op =) head_tfrees (a, S) then I
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   235
      else insert (op =) v | _ => I)) rhs [];
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   236
  in
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   237
    if not (check_head head) then
19103
0eb600479944 dest_def: tuned error msg;
wenzelm
parents: 19071
diff changeset
   238
      err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head])
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   239
    else if not (null lhs_bads) then
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   240
      err ("Bad arguments on lhs: " ^ display_terms lhs_bads)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   241
    else if not (null lhs_dups) then
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   242
      err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   243
    else if not (null rhs_extras) then
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   244
      err ("Extra variables on rhs: " ^ display_terms rhs_extras)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   245
    else if not (null rhs_extrasT) then
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   246
      err ("Extra type variables on rhs: " ^ display_types rhs_extrasT)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   247
    else if exists_subterm (fn t => t aconv head) rhs then
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   248
      err "Entity to be defined occurs on rhs"
19071
fdffd7c40864 dest_def: actually return beta-eta contracted equation;
wenzelm
parents: 18964
diff changeset
   249
    else ((lhs, rhs), fold_rev close_arg args (Term.list_all (eq_vars, (mk_equals (lhs, rhs)))))
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   250
  end;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   251
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   252
(*!!x. c x == t[x] to c == %x. t[x]*)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   253
fun abs_def eq =
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   254
  let
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   255
    val body = Term.strip_all_body eq;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   256
    val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq));
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   257
    val (lhs, rhs) = dest_equals (Term.subst_bounds (vars, body));
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   258
    val (lhs', args) = Term.strip_comb lhs;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   259
    val rhs' = Term.list_abs_free (map Term.dest_Free args, rhs);
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   260
  in (lhs', rhs') end;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   261
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   262
fun mk_cond_defpair As (lhs, rhs) =
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   263
  (case Term.head_of lhs of
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   264
    Const (name, _) =>
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   265
      (NameSpace.base name ^ "_def", list_implies (As, mk_equals (lhs, rhs)))
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   266
  | _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs]));
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   267
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   268
fun mk_defpair lhs_rhs = mk_cond_defpair [] lhs_rhs;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   269
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   270
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   271
(** protected propositions **)
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   272
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   273
val protectC = Const ("prop", propT --> propT);
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   274
fun protect t = protectC $ t;
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   275
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   276
fun unprotect (Const ("prop", _) $ t) = t
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   277
  | unprotect t = raise TERM ("unprotect", [t]);
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   278
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   279
18181
wenzelm
parents: 18029
diff changeset
   280
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
(*** Low-level term operations ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   282
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   283
(*Does t occur in u?  Or is alpha-convertible to u?
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   284
  The term t must contain no loose bound variables*)
16846
bbebc68a7faf occs no longer infix (structure not open);
wenzelm
parents: 16130
diff changeset
   285
fun occs (t, u) = exists_subterm (fn s => t aconv s) u;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   286
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   287
(*Close up a formula over all free variables by quantification*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
fun close_form A =
4443
d55e85d7f0c2 term order stuff moved to term.ML;
wenzelm
parents: 4345
diff changeset
   289
  list_all_free (sort_wrt fst (map dest_Free (term_frees A)), A);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   292
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
(*** Specialized operations for resolution... ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   294
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   295
(*computes t(Bound(n+k-1),...,Bound(n))  *)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   296
fun combound (t, n, k) =
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   297
    if  k>0  then  combound (t,n+1,k-1) $ (Bound n)  else  t;
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   298
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   299
(* ([xn,...,x1], t)   ======>   (x1,...,xn)t *)
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   300
fun rlist_abs ([], body) = body
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   301
  | rlist_abs ((a,T)::pairs, body) = rlist_abs(pairs, Abs(a, T, body));
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   302
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   303
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   304
local exception SAME in
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   305
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   306
fun incrT k =
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   307
  let
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   308
    fun incr (TVar ((a, i), S)) = TVar ((a, i + k), S)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   309
      | incr (Type (a, Ts)) = Type (a, incrs Ts)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   310
      | incr _ = raise SAME
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   311
    and incrs (T :: Ts) =
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   312
        (incr T :: (incrs Ts handle SAME => Ts)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   313
          handle SAME => T :: incrs Ts)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   314
      | incrs [] = raise SAME;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   315
  in incr end;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   316
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
(*For all variables in the term, increment indexnames and lift over the Us
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   318
    result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *)
17120
4ddeef83bd66 optimization to incr_indexes?
paulson
parents: 16879
diff changeset
   319
fun incr_indexes ([], 0) t = t
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   320
  | incr_indexes (Ts, k) t =
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   321
  let
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   322
    val n = length Ts;
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   323
    val incrT = if k = 0 then I else incrT k;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   324
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   325
    fun incr lev (Var ((x, i), T)) =
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   326
          combound (Var ((x, i + k), Ts ---> (incrT T handle SAME => T)), lev, n)
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   327
      | incr lev (Abs (x, T, body)) =
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   328
          (Abs (x, incrT T, incr (lev + 1) body handle SAME => body)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   329
            handle SAME => Abs (x, T, incr (lev + 1) body))
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   330
      | incr lev (t $ u) =
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   331
          (incr lev t $ (incr lev u handle SAME => u)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   332
            handle SAME => t $ incr lev u)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   333
      | incr _ (Const (c, T)) = Const (c, incrT T)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   334
      | incr _ (Free (x, T)) = Free (x, incrT T)
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   335
      | incr _ (t as Bound _) = t;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   336
  in incr 0 t handle SAME => t end;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   337
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   338
fun incr_tvar 0 T = T
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   339
  | incr_tvar k T = incrT k T handle SAME => T;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   340
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   341
end;
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   342
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   343
18248
wenzelm
parents: 18181
diff changeset
   344
(* Lifting functions from subgoal and increment:
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   345
    lift_abs operates on terms
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   346
    lift_all operates on propositions *)
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   347
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   348
fun lift_abs inc =
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   349
  let
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   350
    fun lift Ts (Const ("==>", _) $ _ $ B) t = lift Ts B t
18248
wenzelm
parents: 18181
diff changeset
   351
      | lift Ts (Const ("all", _) $ Abs (a, T, B)) t = Abs (a, T, lift (T :: Ts) B t)
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   352
      | lift Ts _ t = incr_indexes (rev Ts, inc) t;
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   353
  in lift [] end;
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   354
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   355
fun lift_all inc =
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   356
  let
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   357
    fun lift Ts ((c as Const ("==>", _)) $ A $ B) t = c $ A $ lift Ts B t
18248
wenzelm
parents: 18181
diff changeset
   358
      | lift Ts ((c as Const ("all", _)) $ Abs (a, T, B)) t = c $ Abs (a, T, lift (T :: Ts) B t)
18029
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   359
      | lift Ts _ t = incr_indexes (rev Ts, inc) t;
19f1ad18bece renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17120
diff changeset
   360
  in lift [] end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   361
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   362
(*Strips assumptions in goal, yielding list of hypotheses.   *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   363
fun strip_assums_hyp (Const("==>", _) $ H $ B) = H :: strip_assums_hyp B
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
  | strip_assums_hyp (Const("all",_)$Abs(a,T,t)) = strip_assums_hyp t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
  | strip_assums_hyp B = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   367
(*Strips assumptions in goal, yielding conclusion.   *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   368
fun strip_assums_concl (Const("==>", _) $ H $ B) = strip_assums_concl B
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   369
  | strip_assums_concl (Const("all",_)$Abs(a,T,t)) = strip_assums_concl t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   370
  | strip_assums_concl B = B;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
(*Make a list of all the parameters in a subgoal, even if nested*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   373
fun strip_params (Const("==>", _) $ H $ B) = strip_params B
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   374
  | strip_params (Const("all",_)$Abs(a,T,t)) = (a,T) :: strip_params t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   375
  | strip_params B = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
9667
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   377
(*test for meta connectives in prems of a 'subgoal'*)
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   378
fun has_meta_prems prop i =
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   379
  let
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   380
    fun is_meta (Const ("==>", _) $ _ $ _) = true
10442
8ef083987af9 has_meta_prems: include "==";
wenzelm
parents: 9684
diff changeset
   381
      | is_meta (Const ("==", _) $ _ $ _) = true
9667
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   382
      | is_meta (Const ("all", _) $ _) = true
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   383
      | is_meta _ = false;
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   384
  in
13659
3cf622f6b0b2 Removed obsolete functions dealing with flex-flex constraints.
berghofe
parents: 12902
diff changeset
   385
    (case strip_prems (i, [], prop) of
9667
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   386
      (B :: _, _) => exists is_meta (strip_assums_hyp B)
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   387
    | _ => false) handle TERM _ => false
48cefe2daf32 fixed has_meta_prems: strip_assums_hyp;
wenzelm
parents: 9483
diff changeset
   388
  end;
9483
708a8a05497d added has_meta_prems;
wenzelm
parents: 9460
diff changeset
   389
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
(*Removes the parameters from a subgoal and renumber bvars in hypotheses,
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   391
    where j is the total number of parameters (precomputed)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   392
  If n>0 then deletes assumption n. *)
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   393
fun remove_params j n A =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   394
    if j=0 andalso n<=0 then A  (*nothing left to do...*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   395
    else case A of
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   396
        Const("==>", _) $ H $ B =>
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   397
          if n=1 then                           (remove_params j (n-1) B)
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   398
          else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   399
      | Const("all",_)$Abs(a,T,t) => remove_params (j-1) n t
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   400
      | _ => if n>0 then raise TERM("remove_params", [A])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   401
             else A;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   402
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   403
(** Auto-renaming of parameters in subgoals **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   404
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   405
val auto_rename = ref false
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   406
and rename_prefix = ref "ka";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   407
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
(*rename_prefix is not exported; it is set by this function.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   409
fun set_rename_prefix a =
4693
2e47ea2c6109 adapted to baroque chars;
wenzelm
parents: 4679
diff changeset
   410
    if a<>"" andalso forall Symbol.is_letter (Symbol.explode a)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   411
    then  (rename_prefix := a;  auto_rename := true)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   412
    else  error"rename prefix must be nonempty and consist of letters";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   413
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   414
(*Makes parameters in a goal have distinctive names (not guaranteed unique!)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   415
  A name clash could cause the printer to rename bound vars;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   416
    then res_inst_tac would not work properly.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   417
fun rename_vars (a, []) = []
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   418
  | rename_vars (a, (_,T)::vars) =
12902
a23dc0b7566f Symbol.bump_string;
wenzelm
parents: 12796
diff changeset
   419
        (a,T) :: rename_vars (Symbol.bump_string a, vars);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   420
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
(*Move all parameters to the front of the subgoal, renaming them apart;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   422
  if n>0 then deletes assumption n. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
fun flatten_params n A =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   424
    let val params = strip_params A;
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   425
        val vars = if !auto_rename
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   426
                   then rename_vars (!rename_prefix, params)
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   427
                   else ListPair.zip (variantlist(map #1 params,[]),
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   428
                                      map #2 params)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   429
    in  list_all (vars, remove_params (length vars) n A)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   430
    end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   431
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   432
(*Makes parameters in a goal have the names supplied by the list cs.*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   433
fun list_rename_params (cs, Const("==>", _) $ A $ B) =
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   434
      implies $ A $ list_rename_params (cs, B)
9460
53d7ad5bec39 Logic.goal_const;
wenzelm
parents: 5041
diff changeset
   435
  | list_rename_params (c::cs, Const("all",_)$Abs(_,T,t)) =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   436
      all T $ Abs(c, T, list_rename_params (cs, t))
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   437
  | list_rename_params (cs, B) = B;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   438
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   439
(*** Treatmsent of "assume", "erule", etc. ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   440
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   441
(*Strips assumptions in goal yielding
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   442
   HS = [Hn,...,H1],   params = [xm,...,x1], and B,
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   443
  where x1...xm are the parameters. This version (21.1.2005) REQUIRES
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   444
  the the parameters to be flattened, but it allows erule to work on
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   445
  assumptions of the form !!x. phi. Any !! after the outermost string
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   446
  will be regarded as belonging to the conclusion, and left untouched.
15454
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   447
  Used ONLY by assum_pairs.
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   448
      Unless nasms<0, it can terminate the recursion early; that allows
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   449
  erule to work on assumptions of the form P==>Q.*)
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   450
fun strip_assums_imp (0, Hs, B) = (Hs, B)  (*recursion terminated by nasms*)
16879
b81d3f2ee565 incr_tvar (from term.ML), incr_indexes: avoid garbage;
wenzelm
parents: 16862
diff changeset
   451
  | strip_assums_imp (nasms, Hs, Const("==>", _) $ H $ B) =
15454
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   452
      strip_assums_imp (nasms-1, H::Hs, B)
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   453
  | strip_assums_imp (_, Hs, B) = (Hs, B); (*recursion terminated by B*)
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   454
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   455
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   456
(*Strips OUTER parameters only, unlike similar legacy versions.*)
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   457
fun strip_assums_all (params, Const("all",_)$Abs(a,T,t)) =
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   458
      strip_assums_all ((a,T)::params, t)
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   459
  | strip_assums_all (params, B) = (params, B);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   460
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   461
(*Produces disagreement pairs, one for each assumption proof, in order.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   462
  A is the first premise of the lifted rule, and thus has the form
15454
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   463
    H1 ==> ... Hk ==> B   and the pairs are (H1,B),...,(Hk,B).
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   464
  nasms is the number of assumptions in the original subgoal, needed when B
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   465
    has the form B1 ==> B2: it stops B1 from being taken as an assumption. *)
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   466
fun assum_pairs(nasms,A) =
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   467
  let val (params, A') = strip_assums_all ([],A)
15454
4b339d3907a0 thin_tac now works on P==>Q
paulson
parents: 15451
diff changeset
   468
      val (Hs,B) = strip_assums_imp (nasms,[],A')
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   469
      fun abspar t = rlist_abs(params, t)
15451
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   470
      val D = abspar B
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   471
      fun pairrev ([], pairs) = pairs
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   472
        | pairrev (H::Hs, pairs) = pairrev(Hs,  (abspar H, D) :: pairs)
c6c8786b9921 fixed thin_tac with higher-level assumptions by removing the old code to
paulson
parents: 14137
diff changeset
   473
  in  pairrev (Hs,[])
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
  end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   475
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   476
(*Converts Frees to Vars and TFrees to TVars*)
16862
wenzelm
parents: 16846
diff changeset
   477
fun varify (Const(a, T)) = Const (a, Type.varifyT T)
wenzelm
parents: 16846
diff changeset
   478
  | varify (Free (a, T)) = Var ((a, 0), Type.varifyT T)
wenzelm
parents: 16846
diff changeset
   479
  | varify (Var (ixn, T)) = Var (ixn, Type.varifyT T)
wenzelm
parents: 16846
diff changeset
   480
  | varify (t as Bound _) = t
wenzelm
parents: 16846
diff changeset
   481
  | varify (Abs (a, T, body)) = Abs (a, Type.varifyT T, varify body)
wenzelm
parents: 16846
diff changeset
   482
  | varify (f $ t) = varify f $ varify t;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   483
18938
b401ee1cda14 added generic dest_def (mostly from theory.ML);
wenzelm
parents: 18762
diff changeset
   484
(*Inverse of varify.*)
16862
wenzelm
parents: 16846
diff changeset
   485
fun unvarify (Const (a, T)) = Const (a, Type.unvarifyT T)
wenzelm
parents: 16846
diff changeset
   486
  | unvarify (Free (a, T)) = Free (a, Type.unvarifyT T)
wenzelm
parents: 16846
diff changeset
   487
  | unvarify (Var ((a, 0), T)) = Free (a, Type.unvarifyT T)
wenzelm
parents: 16846
diff changeset
   488
  | unvarify (Var (ixn, T)) = Var (ixn, Type.unvarifyT T)  (*non-0 index!*)
wenzelm
parents: 16846
diff changeset
   489
  | unvarify (t as Bound _) = t
wenzelm
parents: 16846
diff changeset
   490
  | unvarify (Abs (a, T, body)) = Abs (a, Type.unvarifyT T, unvarify body)
wenzelm
parents: 16846
diff changeset
   491
  | unvarify (f $ t) = unvarify f $ unvarify t;
546
36e40454e03e /unvarifyT, unvarify: moved to Pure/logic.ML
lcp
parents: 447
diff changeset
   492
13799
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   493
16862
wenzelm
parents: 16846
diff changeset
   494
(* goal states *)
wenzelm
parents: 16846
diff changeset
   495
wenzelm
parents: 16846
diff changeset
   496
fun get_goal st i = nth_prem (i, st)
wenzelm
parents: 16846
diff changeset
   497
  handle TERM _ => error "Goal number out of range";
13799
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   498
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   499
(*reverses parameters for substitution*)
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   500
fun goal_params st i =
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   501
  let val gi = get_goal st i
14137
c57ec95e7763 Removed extraneous rev in function goal_params (the list of parameters
berghofe
parents: 14107
diff changeset
   502
      val rfrees = map Free (rename_wrt_term gi (strip_params gi))
13799
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   503
  in (gi, rfrees) end;
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   504
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   505
fun concl_of_goal st i =
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   506
  let val (gi, rfrees) = goal_params st i
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   507
      val B = strip_assums_concl gi
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   508
  in subst_bounds (rfrees, B) end;
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   509
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   510
fun prems_of_goal st i =
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   511
  let val (gi, rfrees) = goal_params st i
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   512
      val As = strip_assums_hyp gi
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   513
  in map (fn A => subst_bounds (rfrees, A)) As end;
77614fe09362 Moved get_goal, prems_of_goal and concl_of_goal from goals.ML to logic.ML
berghofe
parents: 13659
diff changeset
   514
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   515
end;