src/HOL/Tools/TFL/dcterm.ML
author haftmann
Sat, 28 Aug 2010 16:14:32 +0200
changeset 38864 4abe644fcea5
parent 38795 848be46708dc
child 41164 6854e9a40edc
permissions -rw-r--r--
formerly unnamed infix equality now named HOL.eq
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/dcterm.ML
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     4
    Copyright   1997  University of Cambridge
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
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     8
 * Derived efficient cterm destructors.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     9
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    10
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    11
signature DCTERM =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    12
sig
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    13
  val dest_comb: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    14
  val dest_abs: string option -> cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    15
  val capply: cterm -> cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    16
  val cabs: cterm -> cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    17
  val mk_conj: cterm * cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    18
  val mk_disj: cterm * cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    19
  val mk_exists: cterm * cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    20
  val dest_conj: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    21
  val dest_const: cterm -> {Name: string, Ty: typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    22
  val dest_disj: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    23
  val dest_eq: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    24
  val dest_exists: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    25
  val dest_forall: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    26
  val dest_imp: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    27
  val dest_neg: cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    28
  val dest_pair: cterm -> cterm * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    29
  val dest_var: cterm -> {Name:string, Ty:typ}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    30
  val is_conj: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    31
  val is_disj: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    32
  val is_eq: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    33
  val is_exists: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    34
  val is_forall: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    35
  val is_imp: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    36
  val is_neg: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    37
  val is_pair: cterm -> bool
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    38
  val list_mk_disj: cterm list -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    39
  val strip_abs: cterm -> cterm list * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    40
  val strip_comb: cterm -> cterm * cterm list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    41
  val strip_disj: cterm -> cterm list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    42
  val strip_exists: cterm -> cterm list * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    43
  val strip_forall: cterm -> cterm list * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    44
  val strip_imp: cterm -> cterm list * cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    45
  val drop_prop: cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    46
  val mk_prop: cterm -> cterm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    47
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    48
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    49
structure Dcterm: DCTERM =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    50
struct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    51
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    52
structure U = Utils;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    53
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    54
fun ERR func mesg = U.ERR {module = "Dcterm", func = func, mesg = mesg};
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    55
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    56
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    57
fun dest_comb t = Thm.dest_comb t
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    58
  handle CTERM (msg, _) => raise ERR "dest_comb" msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    59
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    60
fun dest_abs a t = Thm.dest_abs a t
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    61
  handle CTERM (msg, _) => raise ERR "dest_abs" msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    62
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    63
fun capply t u = Thm.capply t u
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    64
  handle CTERM (msg, _) => raise ERR "capply" msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    65
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    66
fun cabs a t = Thm.cabs a t
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    67
  handle CTERM (msg, _) => raise ERR "cabs" msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    68
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    69
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    70
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    71
 * Some simple constructor functions.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    72
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    73
29064
70a61d58460e more antiquotations;
wenzelm
parents: 26626
diff changeset
    74
val mk_hol_const = Thm.cterm_of @{theory HOL} o Const;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    75
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    76
fun mk_exists (r as (Bvar, Body)) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    77
  let val ty = #T(rep_cterm Bvar)
38554
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
    78
      val c = mk_hol_const(@{const_name Ex}, (ty --> HOLogic.boolT) --> HOLogic.boolT)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    79
  in capply c (uncurry cabs r) end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    80
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    81
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
    82
local val c = mk_hol_const(@{const_name HOL.conj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    83
in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    84
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    85
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
    86
local val c = mk_hol_const(@{const_name HOL.disj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    87
in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    88
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    89
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    90
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    91
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    92
 * The primitives.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    93
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    94
fun dest_const ctm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    95
   (case #t(rep_cterm ctm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    96
      of Const(s,ty) => {Name = s, Ty = ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    97
       | _ => raise ERR "dest_const" "not a constant");
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    98
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    99
fun dest_var ctm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   100
   (case #t(rep_cterm ctm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   101
      of Var((s,i),ty) => {Name=s, Ty=ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   102
       | Free(s,ty)    => {Name=s, Ty=ty}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   103
       |             _ => raise ERR "dest_var" "not a variable");
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
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   107
 * Derived destructor operations.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   108
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   109
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   110
fun dest_monop expected tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   111
 let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   112
   fun err () = raise ERR "dest_monop" ("Not a(n) " ^ quote expected);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   113
   val (c, N) = dest_comb tm handle U.ERR _ => err ();
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   114
   val name = #Name (dest_const c handle U.ERR _ => err ());
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   115
 in if name = expected then N else err () end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   116
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   117
fun dest_binop expected tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   118
 let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   119
   fun err () = raise ERR "dest_binop" ("Not a(n) " ^ quote expected);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   120
   val (M, N) = dest_comb tm handle U.ERR _ => err ()
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   121
 in (dest_monop expected M, N) handle U.ERR _ => err () end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   122
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   123
fun dest_binder expected tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   124
  dest_abs NONE (dest_monop expected tm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   125
  handle U.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   126
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   127
38554
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   128
val dest_neg    = dest_monop @{const_name Not}
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   129
val dest_pair   = dest_binop @{const_name Pair}
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   130
val dest_eq     = dest_binop @{const_name HOL.eq}
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38554
diff changeset
   131
val dest_imp    = dest_binop @{const_name HOL.implies}
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   132
val dest_conj   = dest_binop @{const_name HOL.conj}
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   133
val dest_disj   = dest_binop @{const_name HOL.disj}
38554
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   134
val dest_select = dest_binder @{const_name Eps}
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   135
val dest_exists = dest_binder @{const_name Ex}
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   136
val dest_forall = dest_binder @{const_name All}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   137
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   138
(* Query routines *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   139
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   140
val is_eq     = can dest_eq
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   141
val is_imp    = can dest_imp
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   142
val is_select = can dest_select
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   143
val is_forall = can dest_forall
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   144
val is_exists = can dest_exists
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   145
val is_neg    = can dest_neg
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   146
val is_conj   = can dest_conj
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   147
val is_disj   = can dest_disj
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   148
val is_pair   = can dest_pair
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   149
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   151
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   152
 * Iterated creation.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   153
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   154
val list_mk_disj = U.end_itlist (fn d1 => fn tm => mk_disj (d1, tm));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   155
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   156
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   157
 * Iterated destruction. (To the "right" in a term.)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   158
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   159
fun strip break tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   160
  let fun dest (p as (ctm,accum)) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   161
        let val (M,N) = break ctm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   162
        in dest (N, M::accum)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   163
        end handle U.ERR _ => p
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   164
  in dest (tm,[])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   165
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   166
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   167
fun rev2swap (x,l) = (rev l, x);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   168
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   169
val strip_comb   = strip (Library.swap o dest_comb)  (* Goes to the "left" *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   170
val strip_imp    = rev2swap o strip dest_imp
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   171
val strip_abs    = rev2swap o strip (dest_abs NONE)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   172
val strip_forall = rev2swap o strip dest_forall
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   173
val strip_exists = rev2swap o strip dest_exists
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   174
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   175
val strip_disj   = rev o (op::) o strip dest_disj
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   176
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   177
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   178
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   179
 * Going into and out of prop
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   180
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   181
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   182
fun mk_prop ctm =
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 23150
diff changeset
   183
  let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 23150
diff changeset
   184
    val thy = Thm.theory_of_cterm ctm;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 23150
diff changeset
   185
    val t = Thm.term_of ctm;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 23150
diff changeset
   186
  in
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   187
    if can HOLogic.dest_Trueprop t then ctm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   188
    else Thm.cterm_of thy (HOLogic.mk_Trueprop t)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   189
  end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   190
  handle TYPE (msg, _, _) => raise ERR "mk_prop" msg
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   191
    | TERM (msg, _) => raise ERR "mk_prop" msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   192
38554
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 37391
diff changeset
   193
fun drop_prop ctm = dest_monop @{const_name Trueprop} ctm handle U.ERR _ => ctm;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   194
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   195
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   196
end;