src/HOL/Tools/Transfer/transfer_bnf.ML
author wenzelm
Mon, 13 Oct 2014 18:45:48 +0200
changeset 58659 6c9821c32dd5
parent 58458 0c9d59cb3af9
child 58721 f9a966c834bc
permissions -rw-r--r--
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
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 Transfer 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 TRANSFER_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
  val base_name_of_bnf: BNF_Def.bnf -> binding
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    10
  val type_name_of_bnf: BNF_Def.bnf -> string
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    11
  val lookup_defined_pred_data: Proof.context -> string -> Transfer.pred_data
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    12
  val bnf_only_type_ctr: (BNF_Def.bnf -> 'a -> 'a) -> BNF_Def.bnf -> 'a -> 'a
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    13
end
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
structure Transfer_BNF : TRANSFER_BNF =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    16
struct
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    17
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    18
open BNF_Util
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    19
open BNF_Def
56651
fc105315822a localize new size function generation code
blanchet
parents: 56648
diff changeset
    20
open BNF_FP_Util
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    21
open BNF_FP_Def_Sugar
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
(* util functions *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    24
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    25
fun base_name_of_bnf bnf = Binding.name (Binding.name_of (name_of_bnf bnf))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    26
fun mk_Frees_free x Ts ctxt = Variable.variant_frees ctxt [] (mk_names (length Ts) x ~~ Ts) |> map Free
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    27
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    28
fun mk_Domainp P =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    29
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    30
    val PT = fastype_of P
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    31
    val argT = hd (binder_types PT)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    32
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    33
    Const (@{const_name Domainp}, PT --> argT --> HOLogic.boolT) $ P
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    34
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    35
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    36
fun mk_pred pred_def args T =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    37
  let
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    38
    val pred_name = pred_def |> prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    39
      |> head_of |> fst o dest_Const
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    40
    val argsT = map fastype_of args
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    41
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    42
    list_comb (Const (pred_name, argsT ---> (T --> HOLogic.boolT)), args)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    43
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    44
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    45
fun mk_eq_onp arg =
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    46
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    47
    val argT = domain_type (fastype_of arg)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    48
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    49
    Const (@{const_name eq_onp}, (argT --> HOLogic.boolT) --> argT --> argT --> HOLogic.boolT)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    50
      $ arg
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    51
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    52
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    53
fun subst_conv thm =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    54
  Conv.top_sweep_conv (K (Conv.rewr_conv (safe_mk_meta_eq thm))) @{context}
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    55
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    56
fun type_name_of_bnf bnf = T_of_bnf bnf |> dest_Type |> fst
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    57
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    58
fun is_Type (Type _) = true
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    59
  | is_Type _ = false
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    60
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    61
fun bnf_only_type_ctr f bnf = if is_Type (T_of_bnf bnf) then f bnf else I
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
fun bnf_of_fp_sugar (fp_sugar:fp_sugar) = nth (#bnfs (#fp_res fp_sugar)) (#fp_res_index fp_sugar)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    64
56648
2ffa440b3074 manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents: 56538
diff changeset
    65
fun fp_sugar_only_type_ctr f fp_sugars =
2ffa440b3074 manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents: 56538
diff changeset
    66
  (case filter (is_Type o T_of_bnf o bnf_of_fp_sugar) fp_sugars of
2ffa440b3074 manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents: 56538
diff changeset
    67
    [] => I
2ffa440b3074 manual merge + added 'rel_distincts' field to record for symmetry
blanchet
parents: 56538
diff changeset
    68
  | fp_sugars' => f fp_sugars')
56524
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
(* relation constraints - bi_total & co. *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    71
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    72
fun mk_relation_constraint name arg =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    73
  (Const (name, fastype_of arg --> HOLogic.boolT)) $ arg
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    74
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    75
fun side_constraint_tac bnf constr_defs ctxt i =
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    76
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    77
    val thms = constr_defs @ map mk_sym [rel_eq_of_bnf bnf, rel_conversep_of_bnf bnf,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    78
      rel_OO_of_bnf bnf]
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    79
  in
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    80
    (SELECT_GOAL (Local_Defs.unfold_tac ctxt thms) THEN' rtac (rel_mono_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    81
      THEN_ALL_NEW atac) i
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    82
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    83
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    84
fun bi_constraint_tac constr_iff sided_constr_intros ctxt i =
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    85
  (SELECT_GOAL (Local_Defs.unfold_tac ctxt [constr_iff]) THEN'
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    86
    CONJ_WRAP' (fn thm => rtac thm THEN_ALL_NEW (REPEAT_DETERM o etac conjE THEN' atac)) sided_constr_intros) i
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
fun generate_relation_constraint_goal ctxt bnf constraint_def =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    89
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    90
    val constr_name = constraint_def |> prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    91
      |> head_of |> fst o dest_Const
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    92
    val live = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    93
    val (((As, Bs), Ds), ctxt) = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    94
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    95
      ||>> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    96
      ||>> mk_TFrees (dead_of_bnf bnf)
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
    97
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    98
    val relator = mk_rel_of_bnf Ds As Bs bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
    99
    val relsT = map2 mk_pred2T As Bs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   100
    val (args, ctxt) = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   101
    val concl = HOLogic.mk_Trueprop (mk_relation_constraint constr_name (list_comb (relator, args)))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   102
    val assms = map (HOLogic.mk_Trueprop o (mk_relation_constraint constr_name)) args
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   103
    val goal = Logic.list_implies (assms, concl)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   104
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   105
    (goal, ctxt)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   106
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   107
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   108
fun prove_relation_side_constraint ctxt bnf constraint_def =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   109
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   110
    val old_ctxt = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   111
    val (goal, ctxt) = generate_relation_constraint_goal ctxt bnf constraint_def
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 57399
diff changeset
   112
    val thm = Goal.prove_sorry ctxt [] [] goal
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   113
      (fn {context = ctxt, prems = _} => side_constraint_tac bnf [constraint_def] 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: 57399
diff changeset
   114
      |> Thm.close_derivation
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   115
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   116
    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
   117
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   118
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   119
fun prove_relation_bi_constraint ctxt bnf constraint_def side_constraints =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   120
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   121
    val old_ctxt = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   122
    val (goal, ctxt) = generate_relation_constraint_goal ctxt bnf constraint_def
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 57399
diff changeset
   123
    val thm = Goal.prove_sorry ctxt [] [] goal
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   124
      (fn {context = ctxt, prems = _} => bi_constraint_tac constraint_def side_constraints 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: 57399
diff changeset
   125
      |> Thm.close_derivation
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   126
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   127
    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
   128
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   129
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   130
val defs = [("left_total_rel", @{thm left_total_alt_def}), ("right_total_rel", @{thm right_total_alt_def}),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   131
  ("left_unique_rel", @{thm left_unique_alt_def}), ("right_unique_rel", @{thm right_unique_alt_def})]
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   132
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   133
fun prove_relation_constraints bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   134
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   135
    val transfer_attr = @{attributes [transfer_rule]}
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   136
    val Tname = base_name_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   137
    fun qualify suffix = Binding.qualified true suffix Tname
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   138
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   139
    val defs = map (apsnd (prove_relation_side_constraint lthy bnf)) defs
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   140
    val bi_total = prove_relation_bi_constraint lthy bnf @{thm bi_total_alt_def}
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   141
      [snd (nth defs 0), snd (nth defs 1)]
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   142
    val bi_unique = prove_relation_bi_constraint lthy bnf @{thm bi_unique_alt_def}
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   143
      [snd (nth defs 2), snd (nth defs 3)]
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   144
    val defs = ("bi_total_rel", bi_total) :: ("bi_unique_rel", bi_unique) :: defs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   145
    val notes = maps (fn (name, thm) => [((qualify name, []), [([thm], transfer_attr)])]) defs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   146
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   147
    notes
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   148
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   149
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   150
(* relator_eq *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   151
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   152
fun relator_eq bnf =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   153
  [((Binding.empty, []), [([rel_eq_of_bnf bnf], @{attributes [relator_eq]})])]
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   154
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   155
(* predicator definition and Domainp and eq_onp theorem *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   156
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   157
fun define_pred bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   158
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   159
    fun mk_pred_name c = Binding.prefix_name "pred_" c
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   160
    val live = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   161
    val Tname = base_name_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   162
    val ((As, Ds), lthy) = lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   163
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   164
      ||>> mk_TFrees (dead_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   165
    val T = mk_T_of_bnf Ds As bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   166
    val sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   167
    val argTs = map mk_pred1T As
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   168
    val args = mk_Frees_free "P" argTs lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   169
    val conjs = map (fn (set, arg) => mk_Ball (set $ Bound 0) arg) (sets ~~ args)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   170
    val rhs = Abs ("x", T, foldr1 HOLogic.mk_conj conjs)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   171
    val pred_name = mk_pred_name Tname
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   172
    val headT = argTs ---> (T --> HOLogic.boolT)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   173
    val head = Free (Binding.name_of pred_name, headT)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   174
    val lhs = list_comb (head, args)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   175
    val def = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   176
    val ((_, (_, pred_def)), lthy) = Specification.definition ((SOME (pred_name, SOME headT, NoSyn)),
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   177
      ((Binding.empty, []), def)) lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   178
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   179
    (pred_def, lthy)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   180
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   181
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   182
fun Domainp_tac bnf pred_def ctxt i =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   183
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   184
    val n = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   185
    val set_map's = set_map_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   186
  in
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   187
    EVERY' [rtac ext, SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm Domainp.simps},
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   188
        in_rel_of_bnf bnf, pred_def]), rtac iffI,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   189
        REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE], hyp_subst_tac ctxt,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   190
        CONJ_WRAP' (fn set_map => EVERY' [rtac ballI, dtac (set_map RS equalityD1 RS set_mp),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   191
          etac imageE, dtac set_rev_mp, atac, REPEAT_DETERM o eresolve_tac [CollectE, @{thm case_prodE}],
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   192
          hyp_subst_tac ctxt, rtac @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]},
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   193
          etac @{thm DomainPI}]) set_map's,
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   194
        REPEAT_DETERM o etac conjE, REPEAT_DETERM o resolve_tac [exI, (refl RS conjI), rotate_prems 1 conjI],
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   195
        rtac refl, rtac (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   196
          map_id_of_bnf bnf]),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   197
        REPEAT_DETERM_N n o (EVERY' [rtac @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]},
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   198
          rtac @{thm fst_conv}]), rtac CollectI,
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   199
        CONJ_WRAP' (fn set_map => EVERY' [rtac (set_map RS @{thm ord_eq_le_trans}),
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   200
          REPEAT_DETERM o resolve_tac [@{thm image_subsetI}, CollectI, @{thm case_prodI}],
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   201
          dtac (rotate_prems 1 bspec), atac, etac @{thm DomainpE}, etac @{thm someI}]) set_map's
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   202
      ] i
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   203
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   204
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   205
fun prove_Domainp_rel ctxt bnf pred_def =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   206
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   207
    val live = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   208
    val old_ctxt = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   209
    val (((As, Bs), Ds), ctxt) = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   210
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   211
      ||>> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   212
      ||>> mk_TFrees (dead_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   213
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   214
    val relator = mk_rel_of_bnf Ds As Bs bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   215
    val relsT = map2 mk_pred2T As Bs
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   216
    val T = mk_T_of_bnf Ds As bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   217
    val (args, ctxt) = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   218
    val lhs = mk_Domainp (list_comb (relator, args))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   219
    val rhs = mk_pred pred_def (map mk_Domainp args) T
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   220
    val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 57399
diff changeset
   221
    val thm = Goal.prove_sorry ctxt [] [] goal
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   222
      (fn {context = ctxt, prems = _} => Domainp_tac bnf pred_def 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: 57399
diff changeset
   223
      |> Thm.close_derivation
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   224
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   225
    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
   226
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   227
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   228
fun pred_eq_onp_tac bnf pred_def ctxt i =
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   229
  (SELECT_GOAL (Local_Defs.unfold_tac ctxt [@{thm eq_onp_Grp},
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   230
    @{thm Ball_Collect}, pred_def]) THEN' CONVERSION (subst_conv (map_id0_of_bnf bnf RS sym))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   231
  THEN' rtac (rel_Grp_of_bnf bnf)) i
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   232
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   233
fun prove_rel_eq_onp ctxt bnf pred_def =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   234
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   235
    val live = live_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   236
    val old_ctxt = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   237
    val ((As, Ds), ctxt) = ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   238
      |> mk_TFrees live
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   239
      ||>> mk_TFrees (dead_of_bnf bnf)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   240
    val T = mk_T_of_bnf Ds As bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   241
    val argTs = map mk_pred1T As
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   242
    val (args, ctxt) = mk_Frees "P" argTs ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   243
    val relator = mk_rel_of_bnf Ds As As bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   244
    val lhs = list_comb (relator, map mk_eq_onp args)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   245
    val rhs = mk_eq_onp (mk_pred pred_def args T)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   246
    val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop
57890
1e13f63fb452 use Goal.prove_sorry instead of Goal.prove where possible (= no schematics in goal and tactic is expected to succeed)
traytel
parents: 57399
diff changeset
   247
    val thm = Goal.prove_sorry ctxt [] [] goal
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   248
      (fn {context = ctxt, prems = _} => pred_eq_onp_tac bnf pred_def 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: 57399
diff changeset
   249
      |> Thm.close_derivation
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   250
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   251
    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
   252
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   253
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   254
fun predicator bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   255
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   256
    val (pred_def, lthy) = define_pred bnf lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   257
    val pred_def = Morphism.thm (Local_Theory.target_morphism lthy) pred_def
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   258
    val Domainp_rel = prove_Domainp_rel lthy bnf pred_def
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   259
    val rel_eq_onp = prove_rel_eq_onp lthy bnf pred_def
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   260
    fun qualify defname suffix = Binding.qualified true suffix defname
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   261
    val Domainp_rel_thm_name = qualify (base_name_of_bnf bnf) "Domainp_rel"
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   262
    val rel_eq_onp_thm_name = qualify (base_name_of_bnf bnf) "rel_eq_onp"
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   263
    val rel_eq_onp_internal = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   264
      (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]})))
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   265
        rel_eq_onp
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   266
    val pred_data = {rel_eq_onp = rel_eq_onp_internal}
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   267
    val type_name = type_name_of_bnf bnf
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   268
    val relator_domain_attr = @{attributes [relator_domain]}
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   269
    val notes = [((Domainp_rel_thm_name, []), [([Domainp_rel], relator_domain_attr)]),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   270
      ((rel_eq_onp_thm_name, []), [([rel_eq_onp], [])])]
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   271
  in
58356
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   272
    lthy
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   273
    |> Local_Theory.notes notes
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   274
    |> snd
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   275
    |> Local_Theory.declaration {syntax = false, pervasive = true}
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   276
      (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data))
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   277
    |> Local_Theory.restore
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   278
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   279
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   280
(* BNF interpretation *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   281
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   282
val transfer_plugin = Plugin_Name.declare_setup @{binding transfer}
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   283
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   284
fun transfer_bnf_interpretation bnf lthy =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   285
  let
58356
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   286
    val dead = dead_of_bnf bnf
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   287
    val constr_notes = if dead > 0 then [] else prove_relation_constraints bnf lthy
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   288
    val relator_eq_notes = if dead > 0 then [] else relator_eq bnf
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   289
  in
58356
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   290
    lthy
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   291
    |> Local_Theory.notes (constr_notes @ relator_eq_notes)
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   292
    |> snd
2f04f1fd28aa added missing 'restore' in 'transfer' plugin
blanchet
parents: 58234
diff changeset
   293
    |> predicator bnf
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   294
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   295
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   296
val _ =
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   297
  Theory.setup
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   298
    (bnf_interpretation transfer_plugin (bnf_only_type_ctr transfer_bnf_interpretation))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   299
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   300
(* simplification rules for the predicator *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   301
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   302
fun lookup_defined_pred_data lthy name =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   303
  case (Transfer.lookup_pred_data lthy name) of
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   304
    SOME data => data
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   305
    | NONE => (error "lookup_pred_data: something went utterly wrong")
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   306
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   307
fun prove_pred_inject lthy (fp_sugar:fp_sugar) =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   308
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   309
    val involved_types = distinct op= (
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   310
        map type_name_of_bnf (#fp_nesting_bnfs fp_sugar)
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   311
      @ map type_name_of_bnf (#live_nesting_bnfs fp_sugar)
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   312
      @ map type_name_of_bnf (#bnfs (#fp_res fp_sugar)))
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   313
    val eq_onps = map (Transfer.rel_eq_onp o lookup_defined_pred_data lthy) involved_types
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   314
    val live = live_of_bnf (bnf_of_fp_sugar fp_sugar)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   315
    val old_lthy = lthy
58227
d91f7a80f412 use right sort constraints
blanchet
parents: 58191
diff changeset
   316
    val old_As = snd (dest_Type (#T fp_sugar))
58230
61e4c96bf2b6 made code work also in the presence of deads
blanchet
parents: 58227
diff changeset
   317
    val liveness =
61e4c96bf2b6 made code work also in the presence of deads
blanchet
parents: 58227
diff changeset
   318
      BNF_FP_Def_Sugar.liveness_of_fp_bnf (length old_As) (hd (#bnfs (#fp_res fp_sugar)));
58227
d91f7a80f412 use right sort constraints
blanchet
parents: 58191
diff changeset
   319
    val (unsorted_As, lthy) = mk_TFrees live lthy
58234
265aea1e9985 honour sorts in N2M
blanchet
parents: 58230
diff changeset
   320
    val As = map2 (Ctr_Sugar_Util.resort_tfree_or_tvar o snd o Ctr_Sugar_Util.dest_TFree_or_TVar)
58230
61e4c96bf2b6 made code work also in the presence of deads
blanchet
parents: 58227
diff changeset
   321
      (map_filter (uncurry (fn true => SOME | false => K NONE)) (liveness ~~ old_As)) unsorted_As
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   322
    val predTs = map mk_pred1T As
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   323
    val (preds, lthy) = mk_Frees "P" predTs lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   324
    val args = map mk_eq_onp preds
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   325
    val cTs = map (SOME o certifyT lthy) (maps (replicate 2) As)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   326
    val cts = map (SOME o certify lthy) args
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   327
    fun get_rhs thm = thm |> concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   328
    fun is_eqn thm = can get_rhs thm
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   329
    fun rel2pred_massage thm =
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   330
      let
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   331
        val live_step = @{lemma "x = y \<Longrightarrow> (eq_onp P a a \<and> x) = (P a \<and> y)" by (simp only: eq_onp_same_args)}
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   332
        val kill_top1 = @{lemma "(top x \<and> P) = P" by blast}
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   333
        val kill_top2 = @{lemma "(P \<and> top x) = P" by blast}
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   334
        fun pred_eq_onp_conj conjs = List.foldr (fn (_, thm) => thm RS live_step)
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   335
          @{thm refl[of True]} conjs
56538
7c3b6b799b94 observe also DEADID BNFs and associate the conjunction in rel_inject to the right
kuncar
parents: 56524
diff changeset
   336
        val conjuncts = if is_eqn thm then thm |> get_rhs |> HOLogic.dest_conj else []
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   337
        val kill_top = Local_Defs.unfold lthy [kill_top2] #> Local_Defs.unfold lthy [kill_top1]
56538
7c3b6b799b94 observe also DEADID BNFs and associate the conjunction in rel_inject to the right
kuncar
parents: 56524
diff changeset
   338
        val kill_True = Local_Defs.unfold lthy [@{thm HOL.simp_thms(21)}]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   339
      in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   340
        thm
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   341
        |> Drule.instantiate' cTs cts
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   342
        |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   343
          (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]})))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   344
        |> Local_Defs.unfold lthy eq_onps
57399
cfc19f0a6261 compile
blanchet
parents: 56677
diff changeset
   345
        |> (fn thm => if conjuncts <> [] then @{thm box_equals}
56538
7c3b6b799b94 observe also DEADID BNFs and associate the conjunction in rel_inject to the right
kuncar
parents: 56524
diff changeset
   346
              OF [thm, @{thm eq_onp_same_args}, pred_eq_onp_conj conjuncts |> kill_True]
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   347
            else thm RS (@{thm eq_onp_same_args} RS iffD1))
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   348
        |> kill_top
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   349
      end
58458
0c9d59cb3af9 refactor fp_sugar move theorems
desharna
parents: 58356
diff changeset
   350
    val rel_injects = #rel_injects (#fp_bnf_sugar fp_sugar)
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   351
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   352
    rel_injects
56538
7c3b6b799b94 observe also DEADID BNFs and associate the conjunction in rel_inject to the right
kuncar
parents: 56524
diff changeset
   353
    |> map (Local_Defs.unfold lthy [@{thm conj_assoc}])
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   354
    |> map rel2pred_massage
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   355
    |> Variable.export lthy old_lthy
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   356
    |> map Drule.zero_var_indexes
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   357
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   358
56677
660ffb526069 predicator simplification rules: support also partially specialized types e.g. 'a * nat
kuncar
parents: 56651
diff changeset
   359
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   360
(* fp_sugar interpretation *)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   361
58186
a6c3962ea907 named interpretations
blanchet
parents: 58179
diff changeset
   362
fun transfer_fp_sugars_interpretation fp_sugar lthy =
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   363
  let
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   364
    val pred_injects = prove_pred_inject lthy fp_sugar
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   365
    fun qualify defname suffix = Binding.qualified true suffix defname
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   366
    val pred_inject_thm_name = qualify (base_name_of_bnf (bnf_of_fp_sugar fp_sugar)) "pred_inject"
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   367
    val simp_attrs = @{attributes [simp]}
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   368
  in
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   369
    snd (Local_Theory.note ((pred_inject_thm_name, simp_attrs), pred_injects) lthy)
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   370
  end
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   371
58659
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   372
val _ =
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   373
  Theory.setup (fp_sugars_interpretation transfer_plugin
6c9821c32dd5 Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents: 58458
diff changeset
   374
    (fp_sugar_only_type_ctr (fold transfer_fp_sugars_interpretation)))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   375
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents:
diff changeset
   376
end