src/HOL/Tools/Lifting/lifting_def.ML
author kuncar
Tue, 18 Nov 2014 16:19:57 +0100
changeset 60224 e759afe46a8c
parent 60221 45545e6c0984
child 60225 eb4e322734bf
permissions -rw-r--r--
export the result of lifting_def
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Lifting/lifting_def.ML
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     3
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     4
Definitions for constants on quotient types.
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     5
*)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     6
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     7
signature LIFTING_DEF =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     8
sig
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
     9
  type lift_def
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    10
  val rty_of_lift_def: lift_def -> typ
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    11
  val qty_of_lift_def: lift_def -> typ
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    12
  val rhs_of_lift_def: lift_def -> term
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    13
  val lift_const_of_lift_def: lift_def -> term
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    14
  val def_thm_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    15
  val rsp_thm_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    16
  val abs_eq_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    17
  val rep_eq_of_lift_def: lift_def -> thm option
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    18
  val transfer_rules_of_lift_def: lift_def -> thm list
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    19
  val morph_lift_def: morphism -> lift_def -> lift_def
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    20
  val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    21
  val mk_lift_const_of_lift_def: typ -> lift_def -> term
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    22
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    23
  val generate_parametric_transfer_rule:
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    24
    Proof.context -> thm -> thm -> thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    25
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    26
  val add_lift_def:
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    27
    binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
    28
    
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
    29
  val lift_def: 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
    30
    binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> 
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    31
    local_theory -> lift_def * local_theory
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    32
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    33
  val lift_def_cmd:
58011
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57664
diff changeset
    34
    (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list ->
bc6bced136e5 tuned signature -- moved type src to Token, without aliases;
wenzelm
parents: 57664
diff changeset
    35
    local_theory -> Proof.state
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    36
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    37
  val can_generate_code_cert: thm -> bool
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53207
diff changeset
    38
end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    39
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    40
structure Lifting_Def: LIFTING_DEF =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    41
struct
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    42
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47675
diff changeset
    43
open Lifting_Util
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    44
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    45
infix 0 MRSL
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    46
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    47
datatype lift_def = LIFT_DEF of {
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    48
  rty: typ,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    49
  qty: typ,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    50
  rhs: term,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    51
  lift_const: term,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    52
  def_thm: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    53
  rsp_thm: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    54
  abs_eq: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    55
  rep_eq: thm option,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    56
  transfer_rules: thm list
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    57
};
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    58
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    59
fun rep_lift_def (LIFT_DEF lift_def) = lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    60
val rty_of_lift_def = #rty o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    61
val qty_of_lift_def = #qty o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    62
val rhs_of_lift_def = #rhs o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    63
val lift_const_of_lift_def = #lift_const o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    64
val def_thm_of_lift_def = #def_thm o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    65
val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    66
val abs_eq_of_lift_def = #abs_eq o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    67
val rep_eq_of_lift_def = #rep_eq o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    68
val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    69
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    70
fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq transfer_rules =
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    71
  LIFT_DEF {rty = rty, qty = qty,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    72
            rhs = rhs, lift_const = lift_const,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    73
            def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    74
            transfer_rules = transfer_rules };
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    75
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    76
fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    77
  (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const,
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    78
  def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, 
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    79
  transfer_rules = transfer_rules }) =
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    80
  LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const,
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    81
            def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq,
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    82
            transfer_rules = f9 transfer_rules }
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    83
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    84
fun morph_lift_def phi =
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    85
  let
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    86
    val mtyp = Morphism.typ phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    87
    val mterm = Morphism.term phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    88
    val mthm = Morphism.thm phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    89
  in
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    90
    map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) (map mthm)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    91
  end
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    92
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    93
fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    94
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    95
fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    96
  (lift_const_of_lift_def lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    97
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    98
fun inst_of_lift_def ctxt qty lift_def =  mk_inst_of_lift_def qty lift_def
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    99
  |> instT_morphism ctxt |> (fn phi => morph_lift_def phi lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   100
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   101
(* Reflexivity prover *)
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   102
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   103
fun mono_eq_prover ctxt prop =
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   104
  let
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56245
diff changeset
   105
    val refl_rules = Lifting_Info.get_reflexivity_rules ctxt
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56245
diff changeset
   106
    val transfer_rules = Transfer.get_transfer_raw ctxt
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56245
diff changeset
   107
    
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59458
diff changeset
   108
    fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW 
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59458
diff changeset
   109
      (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   110
  in
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   111
    SOME (Goal.prove ctxt [] [] prop (K (main_tac 1)))
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   112
      handle ERROR _ => NONE
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   113
  end
56518
beb3b6851665 left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents: 56245
diff changeset
   114
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   115
fun try_prove_refl_rel ctxt rel =
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   116
  let
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   117
    fun mk_ge_eq x =
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   118
      let
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   119
        val T = fastype_of x
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   120
      in
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   121
        Const (@{const_name "less_eq"}, T --> T --> HOLogic.boolT) $ 
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   122
          (Const (@{const_name HOL.eq}, T)) $ x
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   123
      end;
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   124
    val goal = HOLogic.mk_Trueprop (mk_ge_eq rel);
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   125
  in
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   126
     mono_eq_prover ctxt goal
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   127
  end;
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   128
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   129
fun try_prove_reflexivity ctxt prop =
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   130
  let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   131
    val cprop = Thm.cterm_of ctxt prop
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   132
    val rule = @{thm ge_eq_refl}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   133
    val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   134
    val insts = Thm.first_order_match (concl_pat, cprop)
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   135
    val rule = Drule.instantiate_normalize insts rule
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   136
    val prop = hd (Thm.prems_of rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   137
  in
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   138
    case mono_eq_prover ctxt prop of
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   139
      SOME thm => SOME (thm RS rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   140
      | NONE => NONE
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   141
  end
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   142
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   143
(* 
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   144
  Generates a parametrized transfer rule.
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   145
  transfer_rule - of the form T t f
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   146
  parametric_transfer_rule - of the form par_R t' t
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   147
  
54947
kuncar
parents: 54742
diff changeset
   148
  Result: par_T t' f, after substituing op= for relations in par_R that relate
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   149
    a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   150
    using Lifting_Term.merge_transfer_relations
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   151
*)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   152
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   153
fun generate_parametric_transfer_rule ctxt transfer_rule parametric_transfer_rule =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   154
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   155
    fun preprocess ctxt thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   156
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   157
        val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm;
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   158
        val param_rel = (snd o dest_comb o fst o dest_comb) tm;
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   159
        val free_vars = Term.add_vars param_rel [];
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   160
        
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   161
        fun make_subst (var as (_, typ)) subst = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   162
          let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   163
            val [rty, rty'] = binder_types typ
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   164
          in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   165
            if (Term.is_TVar rty andalso is_Type rty') then
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   166
              (Var var, HOLogic.eq_const rty')::subst
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   167
            else
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   168
              subst
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   169
          end;
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   170
        
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   171
        val subst = fold make_subst free_vars [];
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   172
        val csubst = map (apply2 (Thm.cterm_of ctxt)) subst;
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   173
        val inst_thm = Drule.cterm_instantiate csubst thm;
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   174
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   175
        Conv.fconv_rule 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   176
          ((Conv.concl_conv (Thm.nprems_of inst_thm) o
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   177
            HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54335
diff changeset
   178
            (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   179
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   180
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   181
    fun inst_relcomppI ctxt ant1 ant2 =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   182
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   183
        val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   184
        val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   185
        val fun1 = Thm.cterm_of ctxt (strip_args 2 t1)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   186
        val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   187
        val fun2 = Thm.cterm_of ctxt (strip_args 2 t2)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   188
        val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   189
        val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   190
        val vars = (rev (Term.add_vars (Thm.prop_of relcomppI) []))
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   191
        val subst = map (apfst (Thm.cterm_of ctxt o Var)) (vars ~~ ([fun1] @ args1 @ [fun2] @ args2))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   192
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   193
        Drule.cterm_instantiate subst relcomppI
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   194
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   195
54995
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   196
    fun zip_transfer_rules ctxt thm =
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   197
      let
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   198
        fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   199
        val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   200
        val typ = Thm.typ_of_cterm rel
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   201
        val POS_const = Thm.cterm_of ctxt (mk_POS typ)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   202
        val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ))
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   203
        val goal =
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   204
          Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   205
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   206
        [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply}
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   207
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   208
     
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   209
    val thm =
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   210
      inst_relcomppI ctxt parametric_transfer_rule transfer_rule
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   211
        OF [parametric_transfer_rule, transfer_rule]
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   212
    val preprocessed_thm = preprocess ctxt thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   213
    val orig_ctxt = ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   214
    val (fixed_thm, ctxt) = yield_singleton (apfst snd oo Variable.import true) preprocessed_thm ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   215
    val assms = cprems_of fixed_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   216
    val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add
54995
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   217
    val (prems, ctxt) = fold_map Thm.assume_hyps assms ctxt
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   218
    val ctxt = Context.proof_map (fold add_transfer_rule prems) ctxt
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   219
    val zipped_thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   220
      fixed_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   221
      |> undisch_all
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   222
      |> zip_transfer_rules ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   223
      |> implies_intr_list assms
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   224
      |> singleton (Variable.export ctxt orig_ctxt)
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   225
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   226
    zipped_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   227
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   228
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   229
fun print_generate_transfer_info msg = 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   230
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   231
    val error_msg = cat_lines 
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   232
      ["Generation of a parametric transfer rule failed.",
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   233
      (Pretty.string_of (Pretty.block
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   234
         [Pretty.str "Reason:", Pretty.brk 2, msg]))]
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   235
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   236
    error error_msg
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   237
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   238
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   239
fun map_ter _ x [] = x
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   240
    | map_ter f _ xs = map f xs
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   241
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   242
fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   243
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   244
    val transfer_rule =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   245
      ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer})
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   246
      |> Lifting_Term.parametrize_transfer_rule lthy
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   247
  in
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   248
    (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   249
    handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule]))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   250
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   251
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47675
diff changeset
   252
(* Generation of the code certificate from the rsp theorem *)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   253
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   254
fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   255
  | get_body_types (U, V)  = (U, V)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   256
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   257
fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   258
  | get_binder_types _ = []
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   259
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   260
fun get_binder_types_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = 
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   261
    (T, V) :: get_binder_types_by_rel S (U, W)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   262
  | get_binder_types_by_rel _ _ = []
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   263
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   264
fun get_body_type_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = 
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   265
    get_body_type_by_rel S (U, V)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   266
  | get_body_type_by_rel _ (U, V)  = (U, V)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   267
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   268
fun get_binder_rels (Const (@{const_name "rel_fun"}, _) $ R $ S) = R :: get_binder_rels S
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   269
  | get_binder_rels _ = []
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   270
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   271
fun force_rty_type ctxt rty rhs = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   272
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   273
    val thy = Proof_Context.theory_of ctxt
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   274
    val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   275
    val rty_schematic = fastype_of rhs_schematic
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   276
    val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   277
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   278
    Envir.subst_term_types match rhs_schematic
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   279
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   280
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   281
fun unabs_def ctxt def = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   282
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   283
    val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   284
    fun dest_abs (Abs (var_name, T, _)) = (var_name, T)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   285
      | dest_abs tm = raise TERM("get_abs_var",[tm])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   286
    val (var_name, T) = dest_abs (Thm.term_of rhs)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   287
    val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   288
    val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T)))
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   289
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   290
    Thm.combination def refl_thm |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   291
    singleton (Proof_Context.export ctxt' ctxt)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   292
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   293
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   294
fun unabs_all_def ctxt def = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   295
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   296
    val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   297
    val xs = strip_abs_vars (Thm.term_of rhs)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   298
  in  
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   299
    fold (K (unabs_def ctxt)) xs def
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   300
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   301
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   302
val map_fun_unfolded = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   303
  @{thm map_fun_def[abs_def]} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   304
  unabs_def @{context} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   305
  unabs_def @{context} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   306
  Local_Defs.unfold @{context} [@{thm comp_def}]
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   307
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   308
fun unfold_fun_maps ctm =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   309
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   310
    fun unfold_conv ctm =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   311
      case (Thm.term_of ctm) of
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   312
        Const (@{const_name "map_fun"}, _) $ _ $ _ => 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   313
          (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   314
        | _ => Conv.all_conv ctm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   315
  in
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   316
    (Conv.fun_conv unfold_conv) ctm
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   317
  end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   318
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   319
fun unfold_fun_maps_beta ctm =
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   320
  let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   321
  in 
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   322
    (unfold_fun_maps then_conv try_beta_conv) ctm 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   323
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   324
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   325
fun prove_rel ctxt rsp_thm (rty, qty) =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   326
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   327
    val ty_args = get_binder_types (rty, qty)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   328
    fun disch_arg args_ty thm = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   329
      let
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   330
        val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   331
      in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   332
        [quot_thm, thm] MRSL @{thm apply_rsp''}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   333
      end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   334
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   335
    fold disch_arg ty_args rsp_thm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   336
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   337
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   338
exception CODE_CERT_GEN of string
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   339
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   340
fun simplify_code_eq ctxt def_thm = 
48024
7599b28b707f don't be so aggressive during unfolding id and o
kuncar
parents: 47982
diff changeset
   341
  Local_Defs.unfold ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   342
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   343
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   344
  quot_thm - quotient theorem (Quotient R Abs Rep T).
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   345
  returns: whether the Lifting package is capable to generate code for the abstract type
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   346
    represented by quot_thm
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   347
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   348
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   349
fun can_generate_code_cert quot_thm  =
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   350
  case quot_thm_rel quot_thm of
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   351
    Const (@{const_name HOL.eq}, _) => true
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   352
    | Const (@{const_name eq_onp}, _) $ _  => true
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   353
    | _ => false
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   354
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   355
fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   356
  let
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   357
    val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   358
    val unabs_def = unabs_all_def ctxt unfolded_def
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   359
  in  
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   360
    if body_type rty = body_type qty then 
55723
f66371633e13 be consistent and produce always rep_eq with =
kuncar
parents: 55610
diff changeset
   361
      SOME (simplify_code_eq ctxt (unabs_def RS @{thm meta_eq_to_obj_eq}))
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   362
    else 
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   363
      let
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   364
        val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty))
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   365
        val rel_fun = prove_rel ctxt rsp_thm (rty, qty)
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   366
        val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq}
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   367
      in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   368
        case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   369
          SOME mono_eq_thm =>
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   370
            let
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   371
              val rep_abs_eq = mono_eq_thm RS rep_abs_thm
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   372
              val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm)
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   373
              val rep_refl = Thm.reflexive rep RS @{thm meta_eq_to_obj_eq}
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   374
              val repped_eq = [rep_refl, unabs_def RS @{thm meta_eq_to_obj_eq}] MRSL @{thm cong}
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   375
              val code_cert = [repped_eq, rep_abs_eq] MRSL trans
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   376
            in
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   377
              SOME (simplify_code_eq ctxt code_cert)
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   378
            end
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   379
          | NONE => NONE
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   380
      end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   381
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   382
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   383
fun generate_abs_eq ctxt def_thm rsp_thm quot_thm =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   384
  let
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   385
    val abs_eq_with_assms =
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   386
      let
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   387
        val (rty, qty) = quot_thm_rty_qty quot_thm
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   388
        val rel = quot_thm_rel quot_thm
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   389
        val ty_args = get_binder_types_by_rel rel (rty, qty)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   390
        val body_type = get_body_type_by_rel rel (rty, qty)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   391
        val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   392
        
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   393
        val rep_abs_folded_unmapped_thm = 
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   394
          let
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   395
            val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   396
            val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id)
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   397
            val unfolded_maps_eq = unfold_fun_maps ctm
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   398
            val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap}
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   399
            val prems_pat = (hd o Drule.cprems_of) t1
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   400
            val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq)
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   401
          in
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   402
            unfolded_maps_eq RS (Drule.instantiate_normalize insts t1)
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   403
          end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   404
      in
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   405
        rep_abs_folded_unmapped_thm
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   406
        |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   407
        |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm]))
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   408
      end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   409
    
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   410
    val prem_rels = get_binder_rels (quot_thm_rel quot_thm);
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   411
    val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt) 
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   412
      |> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the) 
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   413
      |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl}))
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   414
    val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   415
  in
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   416
    simplify_code_eq ctxt abs_eq
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   417
  end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   418
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   419
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   420
fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy =
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   421
  let
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   422
    fun no_abstr (t $ u) = no_abstr t andalso no_abstr u
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   423
      | no_abstr (Abs (_, _, t)) = no_abstr t
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   424
      | no_abstr (Const (name, _)) = not (Code.is_abstr thy name)
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   425
      | no_abstr _ = true
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   426
    fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   427
      andalso no_abstr (Thm.prop_of eqn)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   428
    fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   429
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   430
  in
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   431
    if is_valid_eq abs_eq_thm then
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   432
      Code.add_default_eqn abs_eq_thm thy
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   433
    else
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   434
      let
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   435
        val (rty_body, qty_body) = get_body_types (rty, qty)
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   436
      in
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   437
        if rty_body = qty_body then
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   438
         Code.add_default_eqn (the opt_rep_eq_thm) thy
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   439
        else
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   440
          if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm)
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   441
          then
59458
9de8ac92cafa abstract code equation may also be default
haftmann
parents: 59058
diff changeset
   442
            Code.add_abs_default_eqn (the opt_rep_eq_thm) thy
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   443
          else
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   444
            thy
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   445
      end
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   446
  end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   447
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   448
local
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   449
  fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = 
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   450
    let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   451
      fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   452
    in
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   453
      Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty]
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   454
    end
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   455
  
57664
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   456
  exception DECODE
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   457
    
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   458
  fun decode_code_eq thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   459
    if Thm.nprems_of thm > 0 then raise DECODE 
57664
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   460
    else
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   461
      let
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   462
        val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   463
        val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   464
        fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type
57664
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   465
      in
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   466
        (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) 
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   467
      end
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   468
  
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   469
  fun register_encoded_code_eq thm thy =
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   470
    let
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   471
      val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   472
    in
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   473
      register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   474
    end
57664
179b9c43fe84 having extra assumptions (typically from a context) means there is no chance to have a valid code equation => skip decoding and registration of the code equations
kuncar
parents: 57663
diff changeset
   475
    handle DECODE => thy
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   476
  
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   477
  val register_code_eq_attribute = Thm.declaration_attribute
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   478
    (fn thm => Context.mapping (register_encoded_code_eq thm) I)
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   479
  val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute)
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   480
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   481
  fun no_no_code ctxt (rty, qty) =
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   482
    if same_type_constrs (rty, qty) then
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   483
      forall (no_no_code ctxt) (Targs rty ~~ Targs qty)
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   484
    else
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   485
      if is_Type qty then
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   486
        if Lifting_Info.is_no_code_type ctxt (Tname qty) then false
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   487
        else 
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   488
          let 
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   489
            val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty)
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   490
            val (rty's, rtyqs) = (Targs rty', Targs rtyq)
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   491
          in
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   492
            forall (no_no_code ctxt) (rty's ~~ rtyqs)
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   493
          end
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   494
      else
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   495
        true
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   496
in
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   497
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   498
fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy =
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   499
  let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   500
    val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   501
  in
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   502
    if no_no_code lthy (rty, qty) then 
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   503
      (snd oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) lthy
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   504
    else
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   505
      lthy
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   506
  end
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   507
end
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   508
            
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   509
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   510
  Defines an operation on an abstract type in terms of a corresponding operation 
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   511
    on a representation type.
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   512
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   513
  var - a binding and a mixfix of the new constant being defined
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   514
  qty - an abstract type of the new constant
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   515
  rhs - a term representing the new constant on the raw level
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   516
  rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'),
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   517
    i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs"
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   518
  par_thms - a parametricity theorem for rhs
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   519
*)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   520
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   521
fun add_lift_def var qty rhs rsp_thm par_thms lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   522
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   523
    val rty = fastype_of rhs
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   524
    val quot_thm = Lifting_Term.prove_quot_thm lthy (rty, qty)
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   525
    val absrep_trm =  quot_thm_abs quot_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   526
    val rty_forced = (domain_type o fastype_of) absrep_trm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   527
    val forced_rhs = force_rty_type lthy rty_forced rhs
53207
9745b7d4cc79 prefer Binding.name_of over Binding.print -- the latter leads to funny quotes and markup within the constructed term;
wenzelm
parents: 53205
diff changeset
   528
    val lhs = Free (Binding.name_of (#1 var), qty)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   529
    val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   530
    val (_, prop') = Local_Defs.cert_def lthy prop
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   531
    val (_, newrhs) = Local_Defs.abs_def prop'
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   532
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
   533
    val ((lift_const, (_ , def_thm)), lthy') = 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   534
      Local_Theory.define (var, ((Thm.def_binding (#1 var), []), newrhs)) lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   535
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   536
    val transfer_rules = generate_transfer_rules lthy' quot_thm rsp_thm def_thm par_thms
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49885
diff changeset
   537
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   538
    val abs_eq_thm = generate_abs_eq lthy' def_thm rsp_thm quot_thm
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   539
    val opt_rep_eq_thm = generate_rep_eq lthy' def_thm rsp_thm (rty_forced, qty)
47351
0193e663a19e lift_definition command generates transfer rule
huffman
parents: 47308
diff changeset
   540
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47504
diff changeset
   541
    fun qualify defname suffix = Binding.qualified true suffix defname
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   542
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47504
diff changeset
   543
    val lhs_name = (#1 var)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   544
    val rsp_thm_name = qualify lhs_name "rsp"
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   545
    val abs_eq_thm_name = qualify lhs_name "abs_eq"
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   546
    val rep_eq_thm_name = qualify lhs_name "rep_eq"
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   547
    val transfer_rule_name = qualify lhs_name "transfer"
47373
3c31e6f1b3bd lift_definition declares transfer_rule attribute
huffman
parents: 47361
diff changeset
   548
    val transfer_attr = Attrib.internal (K Transfer.transfer_add)
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
   549
    val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm 
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
   550
      opt_rep_eq_thm transfer_rules 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   551
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   552
    lthy'
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   553
      |> (snd oo Local_Theory.note) ((rsp_thm_name, []), [rsp_thm])
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   554
      |> (snd oo Local_Theory.note) ((transfer_rule_name, [transfer_attr]), transfer_rules)
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   555
      |> (snd oo Local_Theory.note) ((abs_eq_thm_name, []), [abs_eq_thm])
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   556
      |> (case opt_rep_eq_thm of 
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   557
            SOME rep_eq_thm => (snd oo Local_Theory.note) ((rep_eq_thm_name, []), [rep_eq_thm])
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   558
            | NONE => I)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   559
      |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty)
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   560
      |> ` (fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   561
      ||> Local_Theory.restore
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   562
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   563
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   564
local
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   565
  val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm @{context})
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   566
    [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, 
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   567
      @{thm pcr_Domainp}]
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   568
in
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   569
fun mk_readable_rsp_thm_eq tm lthy =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   570
  let
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59586
diff changeset
   571
    val ctm = Thm.cterm_of lthy tm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   572
    
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   573
    fun assms_rewr_conv tactic rule ct =
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   574
      let
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   575
        fun prove_extra_assms thm =
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   576
          let
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   577
            val assms = cprems_of thm
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   578
            fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   579
            fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm))
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   580
          in
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   581
            map_interrupt prove assms
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   582
          end
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   583
    
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   584
        fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm)
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   585
        fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm))
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   586
        fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm))
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   587
        val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule;
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   588
        val lhs = lhs_of rule1;
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   589
        val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1;
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   590
        val rule3 =
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   591
          Thm.instantiate (Thm.match (lhs, ct)) rule2
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   592
            handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]);
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   593
        val proved_assms = prove_extra_assms rule3
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   594
      in
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   595
        case proved_assms of
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   596
          SOME proved_assms =>
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   597
            let
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   598
              val rule3 = proved_assms MRSL rule3
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   599
              val rule4 =
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   600
                if lhs_of rule3 aconvc ct then rule3
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   601
                else
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   602
                  let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3)
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   603
                  in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   604
            in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   605
          | NONE => Conv.no_conv ct
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   606
      end
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   607
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   608
    fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules)
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   609
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   610
    fun simp_arrows_conv ctm =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   611
      let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   612
        val unfold_conv = Conv.rewrs_conv 
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   613
          [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, 
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   614
            @{thm rel_fun_eq_onp_rel[THEN eq_reflection]},
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   615
            @{thm rel_fun_eq[THEN eq_reflection]},
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   616
            @{thm rel_fun_eq_rel[THEN eq_reflection]}, 
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   617
            @{thm rel_fun_def[THEN eq_reflection]}]
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   618
        fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   619
        val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: 
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   620
            eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw lthy)
60221
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   621
        val intro_top_rule = @{thm eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   622
        val kill_tops = Transfer.top_sweep_rewr_conv [@{thm eq_onp_top_eq_eq[THEN eq_reflection]}]
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   623
        val eq_onp_assms_tac = (CONVERSION kill_tops THEN' 
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   624
          TRY o REPEAT_ALL_NEW (resolve_tac lthy eq_onp_assms_tac_rules) 
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   625
          THEN_ALL_NEW (DETERM o Transfer.eq_tac lthy)) 1
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   626
        val relator_eq_onp_conv = Conv.bottom_conv
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   627
          (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac
60221
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   628
            (intro_top_rule :: Lifting_Info.get_relator_eq_onp_rules lthy)))) lthy
45545e6c0984 improve handling of predicators in rsp_thm
kuncar
parents: 60219
diff changeset
   629
          then_conv kill_tops
49626
354f35953800 mk_readable_rsp_thm_eq is more robust now
kuncar
parents: 49625
diff changeset
   630
        val relator_eq_conv = Conv.bottom_conv
354f35953800 mk_readable_rsp_thm_eq is more robust now
kuncar
parents: 49625
diff changeset
   631
          (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq lthy)))) lthy
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   632
      in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   633
        case (Thm.term_of ctm) of
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   634
          Const (@{const_name "rel_fun"}, _) $ _ $ _ => 
49626
354f35953800 mk_readable_rsp_thm_eq is more robust now
kuncar
parents: 49625
diff changeset
   635
            (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   636
          | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   637
      end
47937
70375fa2679d generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents: 47852
diff changeset
   638
    
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   639
    val unfold_ret_val_invs = Conv.bottom_conv 
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56519
diff changeset
   640
      (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy
55773
cbeb32e44ff1 hide Lifting.invariant from a user completely
kuncar
parents: 55737
diff changeset
   641
    val unfold_inv_conv = 
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   642
      Conv.top_sweep_conv (K (Conv.rewr_conv @{thm eq_onp_def[THEN eq_reflection]})) lthy
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   643
    val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv simp_arrows_conv)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   644
    val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   645
    val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   646
    val beta_conv = Thm.beta_conversion true
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   647
    val eq_thm = 
55773
cbeb32e44ff1 hide Lifting.invariant from a user completely
kuncar
parents: 55737
diff changeset
   648
      (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs
cbeb32e44ff1 hide Lifting.invariant from a user completely
kuncar
parents: 55737
diff changeset
   649
         then_conv unfold_inv_conv) ctm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   650
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54335
diff changeset
   651
    Object_Logic.rulify lthy (eq_thm RS Drule.equal_elim_rule2)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   652
  end
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55724
diff changeset
   653
end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   654
47607
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   655
fun rename_to_tnames ctxt term =
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   656
  let
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55945
diff changeset
   657
    fun all_typs (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) = T :: all_typs t
47607
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   658
      | all_typs _ = []
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   659
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55945
diff changeset
   660
    fun rename (Const (@{const_name Pure.all}, T1) $ Abs (_, T2, t)) (new_name :: names) = 
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55945
diff changeset
   661
        (Const (@{const_name Pure.all}, T1) $ Abs (new_name, T2, rename t names)) 
47607
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   662
      | rename t _ = t
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   663
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   664
    val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 58028
diff changeset
   665
    val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t)
47607
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   666
  in
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   667
    rename term new_names
5c17ef8feac7 use tnames for bound variables in rsp thms
kuncar
parents: 47566
diff changeset
   668
  end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   669
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   670
(* This is not very cheap way of getting the rules but we have only few active
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   671
  liftings in the current setting *)
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   672
fun get_cr_pcr_eqs ctxt =
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   673
  let
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   674
    fun collect (data : Lifting_Info.quotient) l =
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   675
      if is_some (#pcr_info data) 
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   676
      then ((Thm.symmetric o safe_mk_meta_eq o #pcr_cr_eq o the o #pcr_info) data :: l) 
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   677
      else l
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   678
    val table = Lifting_Info.get_quotients ctxt
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   679
  in
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   680
    Symtab.fold (fn (_, data) => fn l => collect data l) table []
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   681
  end
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   682
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   683
fun prepare_lift_def var qty rhs par_thms lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   684
  let
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   685
    val rsp_rel = Lifting_Term.equiv_relation lthy (fastype_of rhs, qty)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   686
    val rty_forced = (domain_type o fastype_of) rsp_rel;
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   687
    val forced_rhs = force_rty_type lthy rty_forced rhs;
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   688
    val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv 
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   689
      (Raw_Simplifier.rewrite lthy false (get_cr_pcr_eqs lthy)))
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   690
    val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59586
diff changeset
   691
      |> Thm.cterm_of lthy
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   692
      |> cr_to_pcr_conv
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   693
      |> `Thm.concl_of
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   694
      |>> Logic.dest_equals
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   695
      |>> snd
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   696
    val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   697
    val opt_proven_rsp_thm = try_prove_reflexivity lthy prsp_tm
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   698
    
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   699
    fun after_qed internal_rsp_thm lthy = 
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   700
      add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms lthy
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   701
  in
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   702
    case opt_proven_rsp_thm of
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   703
      SOME thm => (NONE, K (after_qed thm))
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   704
      | NONE =>  
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   705
        let
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   706
          val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   707
          val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   708
          val readable_rsp_tm_tnames = rename_to_tnames lthy readable_rsp_tm
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   709
      
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   710
          fun after_qed' thm_list lthy = 
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   711
            let
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   712
              val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm 
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54335
diff changeset
   713
                  (fn {context = ctxt, ...} =>
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54335
diff changeset
   714
                    rtac readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac ctxt (hd thm_list) 1)
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   715
            in
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   716
              after_qed internal_rsp_thm lthy
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   717
            end
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   718
        in
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   719
          (SOME readable_rsp_tm_tnames, after_qed')
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   720
        end 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   721
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   722
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   723
fun lift_def var qty rhs tac par_thms lthy =
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   724
  let
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   725
    val (goal, after_qed) = prepare_lift_def var qty rhs par_thms lthy
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   726
  in
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   727
    case goal of
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   728
      SOME goal => 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   729
        let 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   730
          val rsp_thm = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => tac ctxt)
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   731
            |> Thm.close_derivation
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   732
        in
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   733
          after_qed [[rsp_thm]] lthy
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   734
        end
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   735
      | NONE => after_qed [[Drule.dummy_thm]] lthy
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   736
  end
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   737
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   738
(*
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   739
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   740
  lifting_definition command. It opens a proof of a corresponding respectfulness 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   741
  theorem in a user-friendly, readable form. Then add_lift_def is called internally.
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   742
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   743
*)
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   744
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   745
fun lift_def_cmd (raw_var, rhs_raw, par_xthms) lthy =
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   746
  let
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   747
    val ((binding, SOME qty, mx), lthy) = yield_singleton Proof_Context.read_vars raw_var lthy 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   748
    val var = (binding, mx)
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   749
    val rhs = (Syntax.check_term lthy o Syntax.parse_term lthy) rhs_raw
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   750
    val par_thms = Attrib.eval_thms lthy par_xthms
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   751
    val (goal, after_qed) = prepare_lift_def var qty rhs par_thms lthy
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   752
  in
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
   753
    Proof.theorem NONE (snd oo after_qed) [map (rpair []) (the_list goal)] lthy
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   754
  end
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   755
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   756
fun quot_thm_err ctxt (rty, qty) pretty_msg =
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   757
  let
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   758
    val error_msg = cat_lines
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   759
       ["Lifting failed for the following types:",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   760
        Pretty.string_of (Pretty.block
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   761
         [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]),
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   762
        Pretty.string_of (Pretty.block
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   763
         [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]),
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   764
        "",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   765
        (Pretty.string_of (Pretty.block
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   766
         [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))]
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   767
  in
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   768
    error error_msg
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   769
  end
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   770
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   771
fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) =
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   772
  let
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   773
    val (_, ctxt') = yield_singleton Proof_Context.read_vars raw_var ctxt 
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   774
    val rhs = (Syntax.check_term ctxt' o Syntax.parse_term ctxt') rhs_raw
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   775
    val error_msg = cat_lines
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   776
       ["Lifting failed for the following term:",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   777
        Pretty.string_of (Pretty.block
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   778
         [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]),
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   779
        Pretty.string_of (Pretty.block
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   780
         [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]),
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   781
        "",
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   782
        (Pretty.string_of (Pretty.block
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   783
         [Pretty.str "Reason:", 
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   784
          Pretty.brk 2, 
51957
kuncar
parents: 51927
diff changeset
   785
          Pretty.str "The type of the term cannot be instantiated to",
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   786
          Pretty.brk 1,
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   787
          Pretty.quote (Syntax.pretty_typ ctxt rty_forced),
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   788
          Pretty.str "."]))]
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   789
    in
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   790
      error error_msg
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   791
    end
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   792
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   793
fun lift_def_cmd_with_err_handling (raw_var, rhs_raw, par_xthms) lthy =
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   794
  (lift_def_cmd (raw_var, rhs_raw, par_xthms) lthy
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   795
    handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg)
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   796
    handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => 
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   797
      check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw)
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   798
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   799
(* parser and command *)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   800
val liftdef_parser =
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   801
  (((Parse.binding -- (@{keyword "::"} |-- (Parse.typ >> SOME) -- Parse.opt_mixfix')) >> Parse.triple2)
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   802
    --| @{keyword "is"} -- Parse.term -- 
58028
e4250d370657 tuned signature -- define some elementary operations earlier;
wenzelm
parents: 58011
diff changeset
   803
      Scan.optional (@{keyword "parametric"} |-- Parse.!!! Parse.xthms1) []) >> Parse.triple1
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   804
val _ =
59936
b8ffc3dc9e24 @{command_spec} is superseded by @{command_keyword};
wenzelm
parents: 59630
diff changeset
   805
  Outer_Syntax.local_theory_to_proof @{command_keyword lift_definition}
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   806
    "definition for constants over the quotient type"
47379
075d22b3a32f detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents: 47373
diff changeset
   807
      (liftdef_parser >> lift_def_cmd_with_err_handling)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   808
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   809
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53207
diff changeset
   810
end (* structure *)