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