src/HOL/Tools/Lifting/lifting_bnf.ML
author blanchet
Mon, 08 Sep 2014 23:09:25 +0200
changeset 58243 3aa25f39cd74
parent 58191 b3c71d630777
child 58634 9f10d82e8188
permissions -rw-r--r--
made 'lifting' plugin more robust
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Transfer/transfer_bnf.ML
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
58191
b3c71d630777 pretend code generation is a ctr_sugar plugin
blanchet
parents: 58186
diff changeset
     9
  val lifting_plugin: string
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    10
end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    11
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    12
structure Lifting_BNF : LIFTING_BNF =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    13
struct
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    14
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
58191
b3c71d630777 pretend code generation is a ctr_sugar plugin
blanchet
parents: 58186
diff changeset
    19
val lifting_plugin = "lifting"
58186
a6c3962ea907 named interpretations
blanchet
parents: 58179
diff changeset
    20
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    21
(* Quotient map theorem *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    22
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    23
fun Quotient_tac bnf ctxt i =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    24
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    25
    val rel_Grp = rel_Grp_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    26
    fun get_lhs thm = thm |> concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    27
    val vars = get_lhs rel_Grp |> strip_comb |> snd |> map_filter (try (strip_comb #> snd #> hd))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    28
    val UNIVs = map (fn var => HOLogic.mk_UNIV (var |> dest_Var |> snd |> dest_Type |> snd |> hd)) vars
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    29
    val inst = map2 (curry(pairself (certify ctxt))) vars UNIVs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    30
    val rel_Grp_UNIV_sym = rel_Grp |> Drule.instantiate_normalize ([], inst) 
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    31
      |> Local_Defs.unfold ctxt @{thms subset_UNIV[THEN eqTrueI] UNIV_def[symmetric] simp_thms(21)}
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    32
      |> (fn thm => thm RS sym)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    33
    val rel_mono = rel_mono_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    34
    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
    35
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    36
    EVERY' [SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Quotient_alt_def5}]), 
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    37
      REPEAT_DETERM o (etac conjE), rtac conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt [rel_Grp_UNIV_sym]),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    38
      rtac rel_mono THEN_ALL_NEW atac, rtac conjI, SELECT_GOAL (Local_Defs.unfold_tac ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    39
        [rel_conversep_sym, rel_Grp_UNIV_sym]), rtac rel_mono THEN_ALL_NEW atac,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    40
      SELECT_GOAL (Local_Defs.unfold_tac ctxt [rel_conversep_sym, rel_OO_of_bnf bnf RS sym]),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    41
      hyp_subst_tac ctxt, rtac refl] i
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    42
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    43
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    44
fun mk_Quotient args =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    45
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    46
    val argTs = map fastype_of args
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    47
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    48
    list_comb (Const (@{const_name Quotient}, argTs ---> HOLogic.boolT), args)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    49
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    50
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    51
fun prove_Quotient_map bnf ctxt =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    52
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    53
    val live = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    54
    val old_ctxt = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    55
    val (((As, Bs), Ds), ctxt) = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    56
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    57
      ||>> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    58
      ||>> mk_TFrees (dead_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    59
    val argTss = map2 (fn a => fn b => [mk_pred2T a a, a --> b, b --> a,mk_pred2T a b]) As Bs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    60
    val ((argss, argss'), ctxt) = fold_map2 mk_Frees ["R", "Abs", "Rep", "T"] (transpose argTss) ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    61
      |>> `transpose
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    62
   
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    63
    val assms = map (mk_Quotient #> HOLogic.mk_Trueprop) argss
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    64
    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
    65
    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
    66
    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
    67
    val T_rel = list_comb (mk_rel_of_bnf Ds As Bs bnf, nth argss' 3)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    68
    val concl = mk_Quotient [R_rel, Abs_map, Rep_map, T_rel] |> HOLogic.mk_Trueprop
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    69
    val goal = Logic.list_implies (assms, concl)
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 56677
diff changeset
    70
    val thm = Goal.prove_sorry ctxt [] [] goal 
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    71
      (fn {context = ctxt, prems = _} => Quotient_tac bnf ctxt 1)
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 56677
diff changeset
    72
      |> Thm.close_derivation
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    73
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    74
    Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    75
  end
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
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    78
fun Quotient_map bnf ctxt =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    79
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    80
    val Quotient = prove_Quotient_map bnf ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    81
    fun qualify defname suffix = Binding.qualified true suffix defname
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    82
    val Quotient_thm_name = qualify (base_name_of_bnf bnf) "Quotient"
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    83
    val notes = [((Quotient_thm_name, []), [([Quotient], @{attributes [quot_map]})])]
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    84
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    85
    notes
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    86
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    87
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    88
(* relator_eq_onp  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    89
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    90
fun relator_eq_onp bnf ctxt =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    91
  let
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56524
diff changeset
    92
    val relator_eq_onp_thm = lookup_defined_pred_data ctxt (type_name_of_bnf bnf)
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56524
diff changeset
    93
      |> Transfer.rel_eq_onp |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv 
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56524
diff changeset
    94
          (Raw_Simplifier.rewrite ctxt false @{thms eq_onp_top_eq_eq[THEN eq_reflection]})))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    95
  in
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56524
diff changeset
    96
    [((Binding.empty, []), [([relator_eq_onp_thm], @{attributes [relator_eq_onp]})])]    
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    97
  end
58243
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
    98
  handle ERROR _ => [] (* the "lifting" plugin should work even if "quotient" is disabled *)
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    99
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   100
(* relator_mono  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   101
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   102
fun relator_mono bnf =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   103
  [((Binding.empty, []), [([rel_mono_of_bnf bnf], @{attributes [relator_mono]})])]    
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   104
  
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   105
(* relator_distr  *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   106
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   107
fun relator_distr bnf =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   108
  [((Binding.empty, []), [([rel_OO_of_bnf bnf RS sym], @{attributes [relator_distr]})])]
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
(* interpretation *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   111
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   112
fun lifting_bnf_interpretation bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   113
  if dead_of_bnf bnf > 0 then lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   114
  else
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   115
    let
58243
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
   116
      val notess = [relator_eq_onp bnf lthy, Quotient_map bnf lthy, relator_mono bnf,
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
   117
        relator_distr bnf]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   118
    in
58243
3aa25f39cd74 made 'lifting' plugin more robust
blanchet
parents: 58191
diff changeset
   119
      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
   120
    end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   121
58191
b3c71d630777 pretend code generation is a ctr_sugar plugin
blanchet
parents: 58186
diff changeset
   122
val _ = Theory.setup (bnf_interpretation lifting_plugin
58186
a6c3962ea907 named interpretations
blanchet
parents: 58179
diff changeset
   123
  (bnf_only_type_ctr (BNF_Util.map_local_theory o lifting_bnf_interpretation))
58179
2de7b0313de3 tuned size function generation
blanchet
parents: 58177
diff changeset
   124
  (bnf_only_type_ctr lifting_bnf_interpretation))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   125
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   126
end