src/HOL/Tools/Lifting/lifting_bnf.ML
author wenzelm
Wed, 22 Jan 2025 22:22:19 +0100
changeset 81954 6f2bcdfa9a19
parent 80672 28e8d402a9ba
permissions -rw-r--r--
misc tuning: more concise operations on prems (without change of exceptions); discontinue odd clone Drule.cprems_of (see also 991a3feaf270);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59720
f893472fff31 proper headers;
wenzelm
parents: 59621
diff changeset
     1
(*  Title:      HOL/Tools/Lifting/lifting_bnf.ML
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar, TU Muenchen
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     3
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     4
Setup for Lifting for types that are BNF.
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     5
*)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     6
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     7
signature LIFTING_BNF =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     8
sig
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     9
end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    10
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    11
structure Lifting_BNF : LIFTING_BNF =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    12
struct
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    13
80672
28e8d402a9ba clarified signature;
wenzelm
parents: 80634
diff changeset
    14
open Lifting_Util
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    15
open BNF_Util
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    16
open BNF_Def
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    17
open Transfer_BNF
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    18
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    19
(* Quotient map theorem *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    20
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    21
fun Quotient_tac bnf ctxt i =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    22
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    23
    val rel_Grp = rel_Grp_of_bnf bnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
    24
    fun get_lhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    25
    val vars = get_lhs rel_Grp |> strip_comb |> snd |> map_filter (try (strip_comb #> snd #> hd))
80634
a90ab1ea6458 tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents: 74282
diff changeset
    26
    val UNIVs = map (fn var => HOLogic.mk_UNIV (var |> dest_Var |> snd |> dest_Type_args |> hd)) vars
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60216
diff changeset
    27
    val inst = map dest_Var vars ~~ map (Thm.cterm_of ctxt) UNIVs
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 70494
diff changeset
    28
    val rel_Grp_UNIV_sym = rel_Grp |> Drule.instantiate_normalize (TVars.empty, Vars.make inst) 
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63003
diff changeset
    29
      |> Local_Defs.unfold0 ctxt @{thms subset_UNIV[THEN eqTrueI] UNIV_def[symmetric] simp_thms(21)}
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    30
      |> (fn thm => thm RS sym)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    31
    val rel_mono = rel_mono_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    32
    val rel_conversep_sym = rel_conversep_of_bnf bnf RS sym
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    33
  in
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63003
diff changeset
    34
    EVERY' [SELECT_GOAL (Local_Defs.unfold0_tac ctxt [@{thm Quotient_alt_def5}]), 
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63003
diff changeset
    35
      REPEAT_DETERM o (etac ctxt conjE), rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [rel_Grp_UNIV_sym]),
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63003
diff changeset
    36
      rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt, rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold0_tac ctxt
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60642
diff changeset
    37
        [rel_conversep_sym, rel_Grp_UNIV_sym]), rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt,
63170
eae6549dbea2 tuned proofs, to allow unfold_abs_def;
wenzelm
parents: 63003
diff changeset
    38
      SELECT_GOAL (Local_Defs.unfold0_tac ctxt [rel_conversep_sym, rel_OO_of_bnf bnf RS sym]),
60728
26ffdb966759 {r,e,d,f}tac with proper context in BNF
traytel
parents: 60642
diff changeset
    39
      hyp_subst_tac ctxt, rtac ctxt refl] i
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    40
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    41
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    42
fun prove_Quotient_map bnf ctxt =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    43
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    44
    val live = live_of_bnf bnf
70320
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    45
    val (((As, Bs), Ds), ctxt') = ctxt
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    46
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    47
      ||>> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    48
      ||>> mk_TFrees (dead_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    49
    val argTss = map2 (fn a => fn b => [mk_pred2T a a, a --> b, b --> a,mk_pred2T a b]) As Bs
70320
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    50
    val ((argss, argss'), ctxt'') = ctxt'
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    51
      |> @{fold_map 2} mk_Frees ["R", "Abs", "Rep", "T"] (transpose argTss)
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    52
      |>> `transpose
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    53
   
80672
28e8d402a9ba clarified signature;
wenzelm
parents: 80634
diff changeset
    54
    val assms = argss |> map (fn [rel, abs, rep, cr] =>
28e8d402a9ba clarified signature;
wenzelm
parents: 80634
diff changeset
    55
      HOLogic.mk_Trueprop (mk_Quotient (rel, abs, rep, cr)))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    56
    val R_rel = list_comb (mk_rel_of_bnf Ds As As bnf, nth argss' 0)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    57
    val Abs_map = list_comb (mk_map_of_bnf Ds As Bs bnf, nth argss' 1)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    58
    val Rep_map = list_comb (mk_map_of_bnf Ds Bs As bnf, nth argss' 2)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    59
    val T_rel = list_comb (mk_rel_of_bnf Ds As Bs bnf, nth argss' 3)
80672
28e8d402a9ba clarified signature;
wenzelm
parents: 80634
diff changeset
    60
    val concl = HOLogic.mk_Trueprop (mk_Quotient (R_rel, Abs_map, Rep_map, T_rel))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    61
    val goal = Logic.list_implies (assms, concl)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    62
  in
70320
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    63
    Goal.prove_sorry ctxt'' [] [] goal
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    64
      (fn {context = goal_ctxt, ...} => Quotient_tac bnf goal_ctxt 1)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 70320
diff changeset
    65
    |> Thm.close_derivation \<^here>
70320
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    66
    |> singleton (Variable.export ctxt'' ctxt)
59258a3192bf misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 69593
diff changeset
    67
    |> Drule.zero_var_indexes
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    68
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    69
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    70
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    71
fun Quotient_map bnf ctxt =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    72
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    73
    val Quotient = prove_Quotient_map bnf ctxt
63003
bf5fcc65586b clarified signature;
wenzelm
parents: 62324
diff changeset
    74
    val Quotient_thm_name = Binding.qualify_name true (base_name_of_bnf bnf) "Quotient"
bf5fcc65586b clarified signature;
wenzelm
parents: 62324
diff changeset
    75
  in [((Quotient_thm_name, []), [([Quotient], @{attributes [quot_map]})])] end
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    76
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    77
(* relator_eq_onp  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    78
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 60728
diff changeset
    79
fun relator_eq_onp bnf =
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63170
diff changeset
    80
  [(Binding.empty_atts, [([rel_eq_onp_of_bnf bnf], @{attributes [relator_eq_onp]})])]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    81
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    82
(* relator_mono  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    83
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    84
fun relator_mono bnf =
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63170
diff changeset
    85
  [(Binding.empty_atts, [([rel_mono_of_bnf bnf], @{attributes [relator_mono]})])]    
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    86
  
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    87
(* relator_distr  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    88
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    89
fun relator_distr bnf =
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63170
diff changeset
    90
  [(Binding.empty_atts, [([rel_OO_of_bnf bnf RS sym], @{attributes [relator_distr]})])]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    91
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    92
(* interpretation *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    93
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    94
fun lifting_bnf_interpretation bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    95
  if dead_of_bnf bnf > 0 then lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    96
  else
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    97
    let
62324
ae44f16dcea5 make predicator a first-class bnf citizen
traytel
parents: 60728
diff changeset
    98
      val notess = [relator_eq_onp bnf, Quotient_map bnf lthy, relator_mono bnf,
58243
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
    99
        relator_distr bnf]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   100
    in
58243
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
   101
      lthy |> fold (perhaps o try o (snd oo Local_Theory.notes)) notess
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   102
    end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   103
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63352
diff changeset
   104
val lifting_plugin = Plugin_Name.declare_setup \<^binding>\<open>lifting\<close>
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   105
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   106
val _ =
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   107
  Theory.setup
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58634
diff changeset
   108
    (bnf_interpretation lifting_plugin (bnf_only_type_ctr lifting_bnf_interpretation))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   109
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   110
end