src/HOL/Tools/Lifting/lifting_util.ML
author wenzelm
Mon, 11 Feb 2013 14:39:04 +0100
changeset 51085 d90218288d51
parent 50226 536ab2e82ead
child 51314 eac4bb5adbf9
permissions -rw-r--r--
make WWW_Find work again, now that its ML modules reside within a theory context (cf. bf5b45870110) -- patch by Rafal Kolanski;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Lifting/lifting_util.ML
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     3
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     4
General-purpose functions used by the Lifting package.
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     5
*)
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     6
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     7
signature LIFTING_UTIL =
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     8
sig
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
     9
  val MRSL: thm list * thm -> thm
47699
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    10
  val Trueprop_conv: conv -> conv
50225
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    11
  val option_fold: 'b -> ('a -> 'b) -> 'a option -> 'b
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    12
  val dest_Quotient: term -> term * term * term * term
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    13
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    14
  val quot_thm_rel: thm -> term
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    15
  val quot_thm_abs: thm -> term
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    16
  val quot_thm_rep: thm -> term
50226
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    17
  val quot_thm_crel: thm -> term
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    18
  val quot_thm_rty_qty: thm -> typ * typ
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    19
end;
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    20
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    21
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    22
structure Lifting_Util: LIFTING_UTIL =
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    23
struct
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    24
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    25
infix 0 MRSL
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    26
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    27
fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    28
47699
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    29
fun Trueprop_conv cv ct =
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    30
  (case Thm.term_of ct of
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    31
    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    32
  | _ => raise CTERM ("Trueprop_conv", [ct]))
bb6b147c6531 added useful Trueprop_conv
kuncar
parents: 47698
diff changeset
    33
50225
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    34
fun option_fold a _ NONE = a
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    35
  | option_fold _ f (SOME x) = f x
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    36
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    37
fun dest_Quotient (Const (@{const_name Quotient}, _) $ rel $ abs $ rep $ cr)
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    38
      = (rel, abs, rep, cr)
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    39
  | dest_Quotient t = raise TERM ("dest_Quotient", [t])
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    40
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    41
(*
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    42
  quot_thm_rel, quot_thm_abs, quot_thm_rep and quot_thm_rty_qty - simple functions
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    43
    for destructing quotient theorems (Quotient R Abs Rep T).
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    44
*)
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    45
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    46
fun quot_thm_rel quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    47
  case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    48
    (rel, _, _, _) => rel
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    49
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    50
fun quot_thm_abs quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    51
  case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    52
    (_, abs, _, _) => abs
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    53
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    54
fun quot_thm_rep quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    55
  case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    56
    (_, _, rep, _) => rep
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    57
50226
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    58
fun quot_thm_crel quot_thm =
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    59
  case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    60
    (_, _, _, crel) => crel
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    61
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    62
fun quot_thm_rty_qty quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    63
  let
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    64
    val abs = quot_thm_abs quot_thm
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    65
    val abs_type = fastype_of abs  
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    66
  in
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    67
    (domain_type abs_type, range_type abs_type)
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    68
  end
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    69
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    70
end;