src/HOL/Tools/Lifting/lifting_util.ML
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 80689 b21af525f543
permissions -rw-r--r--
typo
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
80672
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    10
  val mk_Quotient: term * term * term * term -> term
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    11
  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
    12
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    13
  val quot_thm_rel: thm -> term
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    14
  val quot_thm_abs: thm -> term
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    15
  val quot_thm_rep: thm -> term
50226
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    16
  val quot_thm_crel: thm -> term
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    17
  val quot_thm_rty_qty: thm -> typ * typ
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60223
diff changeset
    18
  val Quotient_conv: conv -> conv -> conv -> conv -> conv
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60223
diff changeset
    19
  val Quotient_R_conv: conv -> conv
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    20
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    21
  val undisch: thm -> thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    22
  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
    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
55454
6ea67a791108 Lifting: support a type variable as a raw type
kuncar
parents: 53651
diff changeset
    26
  val Targs: typ -> typ list
6ea67a791108 Lifting: support a type variable as a raw type
kuncar
parents: 53651
diff changeset
    27
  val Tname: typ -> string
51374
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
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55456
diff changeset
    29
  val map_interrupt: ('a -> 'b option) -> 'a list -> 'b list option
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60223
diff changeset
    30
  val conceal_naming_result: (local_theory -> 'a * local_theory) -> local_theory -> 'a * local_theory
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 52883
diff changeset
    31
end
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    32
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    33
structure Lifting_Util: LIFTING_UTIL =
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    34
struct
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    35
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    36
infix 0 MRSL
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    37
18202d3d5832 move MRSL to a separate file
kuncar
parents:
diff changeset
    38
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
    39
80672
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    40
fun mk_Quotient (rel, abs, rep, cr) =
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    41
  let val \<^Type>\<open>fun A B\<close> = fastype_of abs
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    42
  in \<^Const>\<open>Quotient A B for rel abs rep cr\<close> end
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    43
28e8d402a9ba clarified signature;
wenzelm
parents: 70323
diff changeset
    44
fun dest_Quotient \<^Const_>\<open>Quotient _ _ for rel abs rep cr\<close> = (rel, abs, rep, cr)
47777
f29e7dcd7c40 use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents: 47699
diff changeset
    45
  | 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
    46
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    47
val dest_Quotient_thm = dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    48
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    49
val quot_thm_rel = #1 o dest_Quotient_thm
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    50
val quot_thm_abs = #2 o dest_Quotient_thm
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    51
val quot_thm_rep = #3 o dest_Quotient_thm
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    52
val quot_thm_crel = #4 o dest_Quotient_thm
50226
536ab2e82ead quot_thm_crel
kuncar
parents: 50225
diff changeset
    53
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    54
fun quot_thm_rty_qty quot_thm =
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    55
  let val \<^Type>\<open>fun A B\<close> = fastype_of (quot_thm_abs quot_thm)
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    56
  in (A, B) end
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47777
diff changeset
    57
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    58
fun Quotient_conv R_conv Abs_conv Rep_conv T_conv =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    59
  Conv.combination_conv (Conv.combination_conv
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    60
    (Conv.combination_conv (Conv.arg_conv R_conv) Abs_conv) Rep_conv) T_conv
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    61
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    62
fun Quotient_R_conv R_conv =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    63
  Quotient_conv R_conv Conv.all_conv Conv.all_conv Conv.all_conv
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60223
diff changeset
    64
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    65
fun undisch thm =
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    66
  let val assm = Thm.cprem_of thm 1
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    67
  in Thm.implies_elim thm (Thm.assume assm) end
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    68
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 55945
diff changeset
    69
fun undisch_all thm = funpow (Thm.nprems_of thm) undisch thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    70
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    71
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
    72
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    73
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
    74
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    75
fun all_args_conv conv ctm =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    76
  Conv.try_conv (Conv.combination_conv (all_args_conv conv) conv) ctm
55454
6ea67a791108 Lifting: support a type variable as a raw type
kuncar
parents: 53651
diff changeset
    77
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    78
fun Targs T = if is_Type T then dest_Type_args T else []
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    79
fun Tname T = if is_Type T then dest_Type_name T else ""
55454
6ea67a791108 Lifting: support a type variable as a raw type
kuncar
parents: 53651
diff changeset
    80
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    81
fun relation_types typ =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    82
  (case strip_type typ of
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    83
    ([typ1, typ2], \<^Type>\<open>bool\<close>) => (typ1, typ2)
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    84
  | _ => error "relation_types: not a relation")
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    85
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55456
diff changeset
    86
fun map_interrupt f l =
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55456
diff changeset
    87
  let
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55456
diff changeset
    88
    fun map_interrupt' _ [] l = SOME (rev l)
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    89
      | map_interrupt' f (x::xs) l =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    90
          (case f x of
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    91
            NONE => NONE
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    92
          | SOME v => map_interrupt' f xs (v::l))
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    93
  in map_interrupt' f l [] end
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55456
diff changeset
    94
80689
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    95
fun conceal_naming_result f lthy =
b21af525f543 misc tuning;
wenzelm
parents: 80680
diff changeset
    96
  lthy |> Proof_Context.concealed |> f ||> Proof_Context.restore_naming lthy
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60223
diff changeset
    97
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 52883
diff changeset
    98
end