src/HOL/Tools/TFL/usyntax.ML
author haftmann
Thu, 10 Jun 2010 12:24:03 +0200
changeset 37391 476270a6c2dc
parent 37135 636e6d8645d6
child 37591 d3daea901123
permissions -rw-r--r--
tuned quotes, antiquotations and whitespace
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/TFL/usyntax.ML
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Konrad Slind, Cambridge University Computer Laboratory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     3
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     4
Emulation of HOL's abstract syntax functions.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     5
*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     6
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     7
signature USYNTAX =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     8
sig
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     9
  datatype lambda = VAR   of {Name : string, Ty : typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    10
                  | CONST of {Name : string, Ty : typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    11
                  | COMB  of {Rator: term, Rand : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    12
                  | LAMB  of {Bvar : term, Body : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    13
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    14
  val alpha : typ
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    15
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    16
  (* Types *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    17
  val type_vars  : typ -> typ list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    18
  val type_varsl : typ list -> typ list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    19
  val mk_vartype : string -> typ
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    20
  val is_vartype : typ -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    21
  val strip_prod_type : typ -> typ list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    22
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    23
  (* Terms *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    24
  val free_vars_lr : term -> term list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    25
  val type_vars_in_term : term -> typ list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    26
  val dest_term  : term -> lambda
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    27
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    28
  (* Prelogic *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    29
  val inst      : (typ*typ) list -> term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    30
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    31
  (* Construction routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    32
  val mk_abs    :{Bvar  : term, Body : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    33
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    34
  val mk_imp    :{ant : term, conseq :  term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    35
  val mk_select :{Bvar : term, Body : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    36
  val mk_forall :{Bvar : term, Body : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    37
  val mk_exists :{Bvar : term, Body : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    38
  val mk_conj   :{conj1 : term, conj2 : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    39
  val mk_disj   :{disj1 : term, disj2 : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    40
  val mk_pabs   :{varstruct : term, body : term} -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    41
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    42
  (* Destruction routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    43
  val dest_const: term -> {Name : string, Ty : typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    44
  val dest_comb : term -> {Rator : term, Rand : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    45
  val dest_abs  : string list -> term -> {Bvar : term, Body : term} * string list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    46
  val dest_eq     : term -> {lhs : term, rhs : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    47
  val dest_imp    : term -> {ant : term, conseq : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    48
  val dest_forall : term -> {Bvar : term, Body : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    49
  val dest_exists : term -> {Bvar : term, Body : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    50
  val dest_neg    : term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    51
  val dest_conj   : term -> {conj1 : term, conj2 : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    52
  val dest_disj   : term -> {disj1 : term, disj2 : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    53
  val dest_pair   : term -> {fst : term, snd : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    54
  val dest_pabs   : string list -> term -> {varstruct : term, body : term, used : string list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    55
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    56
  val lhs   : term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    57
  val rhs   : term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    58
  val rand  : term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    59
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    60
  (* Query routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    61
  val is_imp    : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    62
  val is_forall : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    63
  val is_exists : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    64
  val is_neg    : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    65
  val is_conj   : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    66
  val is_disj   : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    67
  val is_pair   : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    68
  val is_pabs   : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    69
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    70
  (* Construction of a term from a list of Preterms *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    71
  val list_mk_abs    : (term list * term) -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    72
  val list_mk_imp    : (term list * term) -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    73
  val list_mk_forall : (term list * term) -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    74
  val list_mk_conj   : term list -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    75
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    76
  (* Destructing a term to a list of Preterms *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    77
  val strip_comb     : term -> (term * term list)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    78
  val strip_abs      : term -> (term list * term)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    79
  val strip_imp      : term -> (term list * term)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    80
  val strip_forall   : term -> (term list * term)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    81
  val strip_exists   : term -> (term list * term)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    82
  val strip_disj     : term -> term list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    83
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    84
  (* Miscellaneous *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    85
  val mk_vstruct : typ -> term list -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    86
  val gen_all    : term -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    87
  val find_term  : (term -> bool) -> term -> term option
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    88
  val dest_relation : term -> term * term * term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    89
  val is_WFR : term -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    90
  val ARB : typ -> term
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    91
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    92
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    93
structure USyntax: USYNTAX =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    94
struct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    95
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    96
infix 4 ##;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    97
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    98
fun USYN_ERR func mesg = Utils.ERR {module = "USyntax", func = func, mesg = mesg};
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    99
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   100
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   101
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   102
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   103
 *                            Types
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   104
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   105
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   106
val mk_prim_vartype = TVar;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   107
fun mk_vartype s = mk_prim_vartype ((s, 0), HOLogic.typeS);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   108
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   109
(* But internally, it's useful *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   110
fun dest_vtype (TVar x) = x
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   111
  | dest_vtype _ = raise USYN_ERR "dest_vtype" "not a flexible type variable";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   112
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   113
val is_vartype = can dest_vtype;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   114
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
   115
val type_vars  = map mk_prim_vartype o OldTerm.typ_tvars
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   116
fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   117
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   118
val alpha  = mk_vartype "'a"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   119
val beta   = mk_vartype "'b"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   120
32287
65d5c5b30747 cleaned up abstract tuple operations and named them consistently
haftmann
parents: 29270
diff changeset
   121
val strip_prod_type = HOLogic.flatten_tupleT;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   122
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   123
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   124
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   125
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   126
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   127
 *                              Terms
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   128
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   129
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   130
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   131
(* Free variables, in order of occurrence, from left to right in the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   132
 * syntax tree. *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   133
fun free_vars_lr tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   134
  let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   135
      fun add (t, frees) = case t of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   136
            Free   _ => if (memb t frees) then frees else t::frees
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   137
          | Abs (_,_,body) => add(body,frees)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   138
          | f$t =>  add(t, add(f, frees))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   139
          | _ => frees
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   140
  in rev(add(tm,[]))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   141
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   142
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   143
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   144
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
   145
val type_vars_in_term = map mk_prim_vartype o OldTerm.term_tvars;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   146
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   147
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   148
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   149
(* Prelogic *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   150
fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   151
fun inst theta = subst_vars (map dest_tybinding theta,[])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   152
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   153
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   154
(* Construction routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   155
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   156
fun mk_abs{Bvar as Var((s,_),ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   157
  | mk_abs{Bvar as Free(s,ty),Body}  = Abs(s,ty,abstract_over(Bvar,Body))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   158
  | mk_abs _ = raise USYN_ERR "mk_abs" "Bvar is not a variable";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   159
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   160
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   161
fun mk_imp{ant,conseq} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   162
   let val c = Const("op -->",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   163
   in list_comb(c,[ant,conseq])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   164
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   165
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   166
fun mk_select (r as {Bvar,Body}) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   167
  let val ty = type_of Bvar
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   168
      val c = Const("Hilbert_Choice.Eps",(ty --> HOLogic.boolT) --> ty)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   169
  in list_comb(c,[mk_abs r])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   170
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   171
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   172
fun mk_forall (r as {Bvar,Body}) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   173
  let val ty = type_of Bvar
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   174
      val c = Const("All",(ty --> HOLogic.boolT) --> HOLogic.boolT)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   175
  in list_comb(c,[mk_abs r])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   176
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   177
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   178
fun mk_exists (r as {Bvar,Body}) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   179
  let val ty = type_of Bvar
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   180
      val c = Const("Ex",(ty --> HOLogic.boolT) --> HOLogic.boolT)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   181
  in list_comb(c,[mk_abs r])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   182
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   183
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   184
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   185
fun mk_conj{conj1,conj2} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   186
   let val c = Const("op &",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   187
   in list_comb(c,[conj1,conj2])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   188
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   189
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   190
fun mk_disj{disj1,disj2} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   191
   let val c = Const("op |",HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   192
   in list_comb(c,[disj1,disj2])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   193
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   194
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   195
fun prod_ty ty1 ty2 = HOLogic.mk_prodT (ty1,ty2);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   196
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   197
local
37135
636e6d8645d6 normalized references to constant "split"
haftmann
parents: 32603
diff changeset
   198
fun mk_uncurry (xt, yt, zt) =
636e6d8645d6 normalized references to constant "split"
haftmann
parents: 32603
diff changeset
   199
    Const(@{const_name split}, (xt --> yt --> zt) --> prod_ty xt yt --> zt)
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37135
diff changeset
   200
fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   201
  | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   202
fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   203
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   204
fun mk_pabs{varstruct,body} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   205
 let fun mpa (varstruct, body) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   206
       if is_var varstruct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   207
       then mk_abs {Bvar = varstruct, Body = body}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   208
       else let val {fst, snd} = dest_pair varstruct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   209
            in mk_uncurry (type_of fst, type_of snd, type_of body) $
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   210
               mpa (fst, mpa (snd, body))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   211
            end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   212
 in mpa (varstruct, body) end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   213
 handle TYPE _ => raise USYN_ERR "mk_pabs" "";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   214
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   215
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   216
(* Destruction routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   217
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   218
datatype lambda = VAR   of {Name : string, Ty : typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   219
                | CONST of {Name : string, Ty : typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   220
                | COMB  of {Rator: term, Rand : term}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   221
                | LAMB  of {Bvar : term, Body : term};
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   222
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   223
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   224
fun dest_term(Var((s,i),ty)) = VAR{Name = s, Ty = ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   225
  | dest_term(Free(s,ty))    = VAR{Name = s, Ty = ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   226
  | dest_term(Const(s,ty))   = CONST{Name = s, Ty = ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   227
  | dest_term(M$N)           = COMB{Rator=M,Rand=N}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   228
  | dest_term(Abs(s,ty,M))   = let  val v = Free(s,ty)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   229
                               in LAMB{Bvar = v, Body = Term.betapply (M,v)}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   230
                               end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   231
  | dest_term(Bound _)       = raise USYN_ERR "dest_term" "Bound";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   232
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   233
fun dest_const(Const(s,ty)) = {Name = s, Ty = ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   234
  | dest_const _ = raise USYN_ERR "dest_const" "not a constant";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   235
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   236
fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   237
  | dest_comb _ =  raise USYN_ERR "dest_comb" "not a comb";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   238
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   239
fun dest_abs used (a as Abs(s, ty, M)) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   240
     let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   241
       val s' = Name.variant used s;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   242
       val v = Free(s', ty);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   243
     in ({Bvar = v, Body = Term.betapply (a,v)}, s'::used)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   244
     end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   245
  | dest_abs _ _ =  raise USYN_ERR "dest_abs" "not an abstraction";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   246
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   247
fun dest_eq(Const("op =",_) $ M $ N) = {lhs=M, rhs=N}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   248
  | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   249
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   250
fun dest_imp(Const("op -->",_) $ M $ N) = {ant=M, conseq=N}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   251
  | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   252
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   253
fun dest_forall(Const("All",_) $ (a as Abs _)) = fst (dest_abs [] a)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   254
  | dest_forall _ = raise USYN_ERR "dest_forall" "not a forall";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   255
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   256
fun dest_exists(Const("Ex",_) $ (a as Abs _)) = fst (dest_abs [] a)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   257
  | dest_exists _ = raise USYN_ERR "dest_exists" "not an existential";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   258
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   259
fun dest_neg(Const("not",_) $ M) = M
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   260
  | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   261
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   262
fun dest_conj(Const("op &",_) $ M $ N) = {conj1=M, conj2=N}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   263
  | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   264
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   265
fun dest_disj(Const("op |",_) $ M $ N) = {disj1=M, disj2=N}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   266
  | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   267
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   268
fun mk_pair{fst,snd} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   269
   let val ty1 = type_of fst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   270
       val ty2 = type_of snd
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37135
diff changeset
   271
       val c = Const(@{const_name Pair},ty1 --> ty2 --> prod_ty ty1 ty2)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   272
   in list_comb(c,[fst,snd])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   273
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   274
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37135
diff changeset
   275
fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   276
  | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   277
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   278
37135
636e6d8645d6 normalized references to constant "split"
haftmann
parents: 32603
diff changeset
   279
local  fun ucheck t = (if #Name (dest_const t) = @{const_name split} then t
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   280
                       else raise Match)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   281
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   282
fun dest_pabs used tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   283
   let val ({Bvar,Body}, used') = dest_abs used tm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   284
   in {varstruct = Bvar, body = Body, used = used'}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   285
   end handle Utils.ERR _ =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   286
          let val {Rator,Rand} = dest_comb tm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   287
              val _ = ucheck Rator
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   288
              val {varstruct = lv, body, used = used'} = dest_pabs used Rand
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   289
              val {varstruct = rv, body, used = used''} = dest_pabs used' body
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   290
          in {varstruct = mk_pair {fst = lv, snd = rv}, body = body, used = used''}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   291
          end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   292
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   293
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   294
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   295
val lhs   = #lhs o dest_eq
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   296
val rhs   = #rhs o dest_eq
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   297
val rand  = #Rand o dest_comb
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   298
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   299
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   300
(* Query routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   301
val is_imp    = can dest_imp
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   302
val is_forall = can dest_forall
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   303
val is_exists = can dest_exists
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   304
val is_neg    = can dest_neg
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   305
val is_conj   = can dest_conj
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   306
val is_disj   = can dest_disj
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   307
val is_pair   = can dest_pair
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   308
val is_pabs   = can (dest_pabs [])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   309
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   310
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   311
(* Construction of a cterm from a list of Terms *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   312
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   313
fun list_mk_abs(L,tm) = fold_rev (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   314
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   315
(* These others are almost never used *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   316
fun list_mk_imp(A,c) = fold_rev (fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   317
fun list_mk_forall(V,t) = fold_rev (fn v => fn b => mk_forall{Bvar=v, Body=b})V t;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   318
val list_mk_conj = Utils.end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm})
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   319
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   320
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   321
(* Need to reverse? *)
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 26750
diff changeset
   322
fun gen_all tm = list_mk_forall(OldTerm.term_frees tm, tm);
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   323
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   324
(* Destructing a cterm to a list of Terms *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   325
fun strip_comb tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   326
   let fun dest(M$N, A) = dest(M, N::A)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   327
         | dest x = x
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   328
   in dest(tm,[])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   329
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   330
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   331
fun strip_abs(tm as Abs _) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   332
       let val ({Bvar,Body}, _) = dest_abs [] tm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   333
           val (bvs, core) = strip_abs Body
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   334
       in (Bvar::bvs, core)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   335
       end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   336
  | strip_abs M = ([],M);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   337
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   338
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   339
fun strip_imp fm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   340
   if (is_imp fm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   341
   then let val {ant,conseq} = dest_imp fm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   342
            val (was,wb) = strip_imp conseq
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   343
        in ((ant::was), wb)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   344
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   345
   else ([],fm);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   346
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   347
fun strip_forall fm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   348
   if (is_forall fm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   349
   then let val {Bvar,Body} = dest_forall fm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   350
            val (bvs,core) = strip_forall Body
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   351
        in ((Bvar::bvs), core)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   352
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   353
   else ([],fm);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   354
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   355
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   356
fun strip_exists fm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   357
   if (is_exists fm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   358
   then let val {Bvar, Body} = dest_exists fm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   359
            val (bvs,core) = strip_exists Body
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   360
        in (Bvar::bvs, core)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   361
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   362
   else ([],fm);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   363
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   364
fun strip_disj w =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   365
   if (is_disj w)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   366
   then let val {disj1,disj2} = dest_disj w
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   367
        in (strip_disj disj1@strip_disj disj2)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   368
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   369
   else [w];
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   370
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   371
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   372
(* Miscellaneous *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   373
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   374
fun mk_vstruct ty V =
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37135
diff changeset
   375
  let fun follow_prod_type (Type(@{type_name "*"},[ty1,ty2])) vs =
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   376
              let val (ltm,vs1) = follow_prod_type ty1 vs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   377
                  val (rtm,vs2) = follow_prod_type ty2 vs1
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   378
              in (mk_pair{fst=ltm, snd=rtm}, vs2) end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   379
        | follow_prod_type _ (v::vs) = (v,vs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   380
  in #1 (follow_prod_type ty V)  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   381
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   382
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   383
(* Search a term for a sub-term satisfying the predicate p. *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   384
fun find_term p =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   385
   let fun find tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   386
      if (p tm) then SOME tm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   387
      else case tm of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   388
          Abs(_,_,body) => find body
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   389
        | (t$u)         => (case find t of NONE => find u | some => some)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   390
        | _             => NONE
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   391
   in find
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   392
   end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   393
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   394
fun dest_relation tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   395
   if (type_of tm = HOLogic.boolT)
37391
476270a6c2dc tuned quotes, antiquotations and whitespace
haftmann
parents: 37135
diff changeset
   396
   then let val (Const("op :",_) $ (Const(@{const_name Pair},_)$y$x) $ R) = tm
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   397
        in (R,y,x)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   398
        end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   399
   else raise USYN_ERR "dest_relation" "not a boolean term";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   400
32603
e08fdd615333 tuned const_name antiquotations
haftmann
parents: 32287
diff changeset
   401
fun is_WFR (Const(@{const_name Wellfounded.wf},_)$_) = true
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   402
  | is_WFR _                 = false;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   403
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   404
fun ARB ty = mk_select{Bvar=Free("v",ty),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   405
                       Body=Const("True",HOLogic.boolT)};
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   406
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   407
end;