src/HOL/Tools/Lifting/lifting_util.ML
author wenzelm
Mon, 05 Aug 2013 15:03:52 +0200
changeset 52861 e93d73b51fd0
parent 51374 84d01fd733cf
child 52883 0a7c97c76f46
permissions -rw-r--r--
commands with overlay remain visible, to avoid loosing printed output;
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
50225
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    10
  val option_fold: 'b -> ('a -> 'b) -> 'a option -> 'b
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    11
  val map_snd: ('b -> 'c) -> ('a * 'b) list -> ('a * 'c) list
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
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    19
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    20
  val undisch: thm -> thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    21
  val undisch_all: thm -> thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    22
  val is_fun_type: typ -> bool
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    23
  val get_args: int -> term -> term list
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    24
  val strip_args: int -> term -> term
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    25
  val all_args_conv: conv -> conv
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    26
  val is_Type: typ -> bool
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    27
  val is_fun_rel: term -> bool
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    28
  val relation_types: typ -> typ * typ
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    29
  val bottom_rewr_conv: thm list -> conv
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    30
  val mk_HOL_eq: thm -> thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    31
  val safe_HOL_meta_eq: thm -> thm
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    32
end;
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    33
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    34
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    35
structure Lifting_Util: LIFTING_UTIL =
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    36
struct
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    37
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    38
infix 0 MRSL
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    39
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    40
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
    41
50225
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    42
fun option_fold a _ NONE = a
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    43
  | option_fold _ f (SOME x) = f x
f478f4ca7f19 add option_fold
kuncar
parents: 47951
diff changeset
    44
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    45
fun map_snd f xs = map (fn (a, b) => (a, f b)) xs
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    46
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    47
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
    48
      = (rel, abs, rep, cr)
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    49
  | 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
    50
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    51
(*
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    52
  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
    53
    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
    54
*)
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    55
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    56
fun quot_thm_rel quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    57
  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
    58
    (rel, _, _, _) => rel
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    59
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    60
fun quot_thm_abs quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    61
  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
    62
    (_, abs, _, _) => abs
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    63
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    64
fun quot_thm_rep quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    65
  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
    66
    (_, _, rep, _) => rep
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    67
50226
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    68
fun quot_thm_crel quot_thm =
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    69
  case (dest_Quotient o HOLogic.dest_Trueprop o prop_of) quot_thm of
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    70
    (_, _, _, crel) => crel
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    71
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    72
fun quot_thm_rty_qty quot_thm =
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    73
  let
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    74
    val abs = quot_thm_abs quot_thm
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    75
    val abs_type = fastype_of abs  
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    76
  in
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    77
    (domain_type abs_type, range_type abs_type)
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    78
  end
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    79
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    80
fun undisch thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    81
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    82
    val assm = Thm.cprem_of thm 1
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    83
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    84
    Thm.implies_elim thm (Thm.assume assm)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    85
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    86
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    87
fun undisch_all thm = funpow (nprems_of thm) undisch thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    88
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    89
fun is_fun_type (Type (@{type_name fun}, _)) = true
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    90
  | is_fun_type _ = false
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    91
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    92
fun get_args n = rev o fst o funpow_yield n (swap o dest_comb)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    93
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    94
fun strip_args n = funpow n (fst o dest_comb)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    95
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    96
fun all_args_conv conv ctm = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    97
  (Conv.combination_conv (Conv.try_conv (all_args_conv conv)) conv) ctm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    98
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    99
fun is_Type (Type _) = true
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   100
  | is_Type _ = false
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   101
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   102
fun is_fun_rel (Const (@{const_name "fun_rel"}, _) $ _ $ _) = true
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   103
  | is_fun_rel _ = false
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   104
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   105
fun relation_types typ = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   106
  case strip_type typ of
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   107
    ([typ1, typ2], @{typ bool}) => (typ1, typ2)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   108
    | _ => error "relation_types: not a relation"
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   109
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   110
fun bottom_rewr_conv rewrs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) @{context}
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   111
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   112
fun mk_HOL_eq r = r RS @{thm meta_eq_to_obj_eq}
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   113
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   114
fun safe_HOL_meta_eq r = mk_HOL_eq r handle Thm.THM _ => r
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   115
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
   116
end;