src/HOL/Tools/Lifting/lifting_def.ML
author wenzelm
Sun, 26 Jul 2015 17:24:54 +0200
changeset 60784 4f590c08fd5d
parent 60239 755e11e2e15d
child 62514 aae510e9a698
permissions -rw-r--r--
updated to infer_instantiate;
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
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
     9
  datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    10
  type lift_def
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    11
  val rty_of_lift_def: lift_def -> typ
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    12
  val qty_of_lift_def: lift_def -> typ
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    13
  val rhs_of_lift_def: lift_def -> term
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    14
  val lift_const_of_lift_def: lift_def -> term
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    15
  val def_thm_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    16
  val rsp_thm_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    17
  val abs_eq_of_lift_def: lift_def -> thm
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    18
  val rep_eq_of_lift_def: lift_def -> thm option
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    19
  val code_eq_of_lift_def: lift_def -> code_eq
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    20
  val transfer_rules_of_lift_def: lift_def -> thm list
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    21
  val morph_lift_def: morphism -> lift_def -> lift_def
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    22
  val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    23
  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
    24
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    25
  type config = { notes: bool }
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60227
diff changeset
    26
  val map_config: (bool -> bool) -> config -> config
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    27
  val default_config: config
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    28
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
    29
  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
    30
    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
    31
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    32
  val add_lift_def: 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    33
    config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    34
      lift_def * local_theory
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    35
  
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    36
  val prepare_lift_def:
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    37
    (binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context -> 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    38
      lift_def * local_theory) -> 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    39
    binding * mixfix -> typ -> term -> thm list -> local_theory -> 
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
    40
    term option * (thm -> Proof.context -> lift_def * local_theory)
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    41
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    42
  val gen_lift_def:
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    43
    (binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    44
      lift_def * local_theory) -> 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    45
    binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    46
    local_theory -> lift_def * local_theory
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    47
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
    48
  val lift_def: 
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
    49
    config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> 
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    50
    local_theory -> lift_def * local_theory
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    51
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    52
  val can_generate_code_cert: thm -> bool
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53207
diff changeset
    53
end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    54
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    55
structure Lifting_Def: LIFTING_DEF =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    56
struct
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    57
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47675
diff changeset
    58
open Lifting_Util
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    59
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    60
infix 0 MRSL
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    61
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
    62
datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    63
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    64
datatype lift_def = LIFT_DEF of {
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    65
  rty: typ,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    66
  qty: typ,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    67
  rhs: term,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    68
  lift_const: term,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    69
  def_thm: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    70
  rsp_thm: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    71
  abs_eq: thm,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    72
  rep_eq: thm option,
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    73
  code_eq: code_eq,
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    74
  transfer_rules: thm list
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    75
};
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    76
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    77
fun rep_lift_def (LIFT_DEF lift_def) = lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    78
val rty_of_lift_def = #rty o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    79
val qty_of_lift_def = #qty o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    80
val rhs_of_lift_def = #rhs o rep_lift_def;
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    81
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
    82
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
    83
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
    84
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
    85
val rep_eq_of_lift_def = #rep_eq o rep_lift_def;
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    86
val code_eq_of_lift_def = #code_eq o rep_lift_def;
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    87
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
    88
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    89
fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules =
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    90
  LIFT_DEF {rty = rty, qty = qty,
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    91
            rhs = rhs, lift_const = lift_const,
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    92
            def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, 
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    93
            code_eq = code_eq, transfer_rules = transfer_rules };
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
    94
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    95
fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    96
  (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const,
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
    97
  def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq,
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    98
  transfer_rules = transfer_rules }) =
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
    99
  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
   100
            def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq,
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   101
            code_eq = f9 code_eq, transfer_rules = f10 transfer_rules }
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   102
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   103
fun morph_lift_def phi =
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   104
  let
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   105
    val mtyp = Morphism.typ phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   106
    val mterm = Morphism.term phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   107
    val mthm = Morphism.thm phi
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   108
  in
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   109
    map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm)
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   110
  end
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   111
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   112
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
   113
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   114
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
   115
  (lift_const_of_lift_def lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   116
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   117
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
   118
  |> instT_morphism ctxt |> (fn phi => morph_lift_def phi lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   119
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   120
(* Config *)
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   121
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   122
type config = { notes: bool };
60229
4cd6462c1fda Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents: 60227
diff changeset
   123
fun map_config f1 { notes = notes } = { notes = f1 notes }
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   124
val default_config = { notes = true };
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   125
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   126
(* Reflexivity prover *)
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   127
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   128
fun mono_eq_prover ctxt prop =
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   129
  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
   130
    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
   131
    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
   132
    
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59458
diff changeset
   133
    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
   134
      (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   135
  in
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   136
    SOME (Goal.prove ctxt [] [] prop (K (main_tac 1)))
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   137
      handle ERROR _ => NONE
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   138
  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
   139
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   140
fun try_prove_refl_rel ctxt rel =
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   141
  let
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   142
    fun mk_ge_eq x =
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   143
      let
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   144
        val T = fastype_of x
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   145
      in
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   146
        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
   147
          (Const (@{const_name HOL.eq}, T)) $ x
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   148
      end;
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   149
    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
   150
  in
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   151
     mono_eq_prover ctxt goal
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   152
  end;
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   153
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   154
fun try_prove_reflexivity ctxt prop =
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   155
  let
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   156
    val cprop = Thm.cterm_of ctxt prop
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   157
    val rule = @{thm ge_eq_refl}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   158
    val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   159
    val insts = Thm.first_order_match (concl_pat, cprop)
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   160
    val rule = Drule.instantiate_normalize insts rule
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   161
    val prop = hd (Thm.prems_of rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   162
  in
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   163
    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
   164
      SOME thm => SOME (thm RS rule)
55609
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   165
      | NONE => NONE
69ac773a467f refactoring
kuncar
parents: 55608
diff changeset
   166
  end
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   167
54335
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   168
(* 
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   169
  Generates a parametrized transfer rule.
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   170
  transfer_rule - of the form T t f
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   171
  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
   172
  
54947
kuncar
parents: 54742
diff changeset
   173
  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
   174
    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
   175
    using Lifting_Term.merge_transfer_relations
03b10317ba78 update documentation of important public ML functions in Lifting
kuncar
parents: 53951
diff changeset
   176
*)
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   177
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   178
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
   179
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   180
    fun preprocess ctxt thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   181
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   182
        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
   183
        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
   184
        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
   185
        
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   186
        fun make_subst (xi, typ) subst = 
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   187
          let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   188
            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
   189
          in
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   190
            if Term.is_TVar rty andalso is_Type rty' then
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   191
              (xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   192
            else
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   193
              subst
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   194
          end;
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   195
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   196
        val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm;
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   197
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   198
        Conv.fconv_rule 
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   199
          ((Conv.concl_conv (Thm.nprems_of inst_thm) o
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   200
            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
   201
            (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
   202
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   203
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   204
    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
   205
      let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   206
        val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   207
        val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   208
        val fun1 = Thm.cterm_of ctxt (strip_args 2 t1)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   209
        val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   210
        val fun2 = Thm.cterm_of ctxt (strip_args 2 t2)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   211
        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
   212
        val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI}
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   213
        val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) []))
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   214
      in
60784
4f590c08fd5d updated to infer_instantiate;
wenzelm
parents: 60239
diff changeset
   215
        infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   216
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   217
54995
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   218
    fun zip_transfer_rules ctxt thm =
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   219
      let
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   220
        fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   221
        val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
   222
        val typ = Thm.typ_of_cterm rel
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   223
        val POS_const = Thm.cterm_of ctxt (mk_POS typ)
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   224
        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
   225
        val goal =
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   226
          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
   227
      in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   228
        [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
   229
      end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   230
     
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   231
    val thm =
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   232
      inst_relcomppI ctxt parametric_transfer_rule transfer_rule
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   233
        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
   234
    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
   235
    val orig_ctxt = ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   236
    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
   237
    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
   238
    val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add
54995
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   239
    val (prems, ctxt) = fold_map Thm.assume_hyps assms ctxt
68228c162ed5 more accurate context;
wenzelm
parents: 54947
diff changeset
   240
    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
   241
    val zipped_thm =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   242
      fixed_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   243
      |> undisch_all
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   244
      |> zip_transfer_rules ctxt
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   245
      |> implies_intr_list assms
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   246
      |> 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
   247
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   248
    zipped_thm
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   249
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   250
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   251
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
   252
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   253
    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
   254
      ["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
   255
      (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
   256
         [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
   257
  in
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   258
    error error_msg
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   259
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   260
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   261
fun map_ter _ x [] = x
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   262
    | map_ter f _ xs = map f xs
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   263
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   264
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
   265
  let
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   266
    val transfer_rule =
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   267
      ([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
   268
      |> 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
   269
  in
53951
03b74ef6d7c6 allow to specify multiple parametricity transfer rules in lift_definition
kuncar
parents: 53651
diff changeset
   270
    (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
   271
    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
   272
  end
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   273
47698
18202d3d5832 move MRSL to a separate file
kuncar
parents: 47675
diff changeset
   274
(* Generation of the code certificate from the rsp theorem *)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   275
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   276
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
   277
  | get_body_types (U, V)  = (U, V)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   278
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   279
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
   280
  | get_binder_types _ = []
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   281
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   282
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
   283
    (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
   284
  | 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
   285
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   286
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
   287
    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
   288
  | 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
   289
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   290
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
   291
  | get_binder_rels _ = []
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   292
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   293
fun force_rty_type ctxt rty rhs = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   294
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   295
    val thy = Proof_Context.theory_of ctxt
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   296
    val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   297
    val rty_schematic = fastype_of rhs_schematic
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   298
    val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   299
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   300
    Envir.subst_term_types match rhs_schematic
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   301
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   302
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   303
fun unabs_def ctxt def = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   304
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   305
    val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   306
    fun dest_abs (Abs (var_name, T, _)) = (var_name, T)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   307
      | dest_abs tm = raise TERM("get_abs_var",[tm])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   308
    val (var_name, T) = dest_abs (Thm.term_of rhs)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   309
    val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   310
    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
   311
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   312
    Thm.combination def refl_thm |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   313
    singleton (Proof_Context.export ctxt' ctxt)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   314
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   315
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   316
fun unabs_all_def ctxt def = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   317
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   318
    val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   319
    val xs = strip_abs_vars (Thm.term_of rhs)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   320
  in  
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   321
    fold (K (unabs_def ctxt)) xs def
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   322
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   323
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   324
val map_fun_unfolded = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   325
  @{thm map_fun_def[abs_def]} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   326
  unabs_def @{context} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   327
  unabs_def @{context} |>
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   328
  Local_Defs.unfold @{context} [@{thm comp_def}]
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   329
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   330
fun unfold_fun_maps ctm =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   331
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   332
    fun unfold_conv ctm =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   333
      case (Thm.term_of ctm) of
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   334
        Const (@{const_name "map_fun"}, _) $ _ $ _ => 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   335
          (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   336
        | _ => Conv.all_conv ctm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   337
  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
   338
    (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
   339
  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
   340
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
   341
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
   342
  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
   343
  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
   344
    (unfold_fun_maps then_conv try_beta_conv) ctm 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   345
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   346
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   347
fun prove_rel ctxt rsp_thm (rty, qty) =
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   348
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   349
    val ty_args = get_binder_types (rty, qty)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   350
    fun disch_arg args_ty thm = 
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   351
      let
47504
aa1b8a59017f go back to the explicit compisition of quotient theorems
kuncar
parents: 47503
diff changeset
   352
        val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   353
      in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   354
        [quot_thm, thm] MRSL @{thm apply_rsp''}
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   355
      end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   356
  in
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   357
    fold disch_arg ty_args rsp_thm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   358
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   359
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   360
exception CODE_CERT_GEN of string
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   361
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   362
fun simplify_code_eq ctxt def_thm = 
48024
7599b28b707f don't be so aggressive during unfolding id and o
kuncar
parents: 47982
diff changeset
   363
  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
   364
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   365
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   366
  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
   367
  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
   368
    represented by quot_thm
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   369
*)
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   370
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   371
fun can_generate_code_cert quot_thm  =
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   372
  case quot_thm_rel quot_thm of
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   373
    Const (@{const_name HOL.eq}, _) => true
56519
c1048f5bbb45 more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents: 56518
diff changeset
   374
    | Const (@{const_name eq_onp}, _) $ _  => true
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   375
    | _ => false
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   376
55610
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   377
fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   378
  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
   379
    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
   380
    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
   381
  in  
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   382
    if body_type rty = body_type qty then 
55723
f66371633e13 be consistent and produce always rep_eq with =
kuncar
parents: 55610
diff changeset
   383
      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
   384
    else 
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   385
      let
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   386
        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
   387
        val rel_fun = prove_rel ctxt rsp_thm (rty, qty)
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   388
        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
   389
      in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   390
        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
   391
          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
   392
            let
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   393
              val rep_abs_eq = mono_eq_thm RS rep_abs_thm
59630
c9aa1c90796f clarified context;
wenzelm
parents: 59621
diff changeset
   394
              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
   395
              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
   396
              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
   397
              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
   398
            in
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   399
              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
   400
            end
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   401
          | NONE => NONE
9066b603dff6 refactoring; generate rep_eq always, not only when it would be accepted by the code generator
kuncar
parents: 55609
diff changeset
   402
      end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   403
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   404
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
   405
fun generate_abs_eq ctxt def_thm rsp_thm quot_thm =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   406
  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
   407
    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
   408
      let
47951
8c8a03765de7 quot_del attribute, it allows us to deregister quotient types
kuncar
parents: 47937
diff changeset
   409
        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
   410
        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
   411
        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
   412
        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
   413
        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
   414
        
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
   415
        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
   416
          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
   417
            val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   418
            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
   419
            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
   420
            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
   421
            val prems_pat = (hd o Drule.cprems_of) t1
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   422
            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
   423
          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
   424
            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
   425
          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
   426
      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
   427
        rep_abs_folded_unmapped_thm
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55773
diff changeset
   428
        |> 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
   429
        |> (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
   430
      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
   431
    
57642
5bc43a73d768 prevent beta-contraction in proving extra assumptions for abs_eq
kuncar
parents: 56540
diff changeset
   432
    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
   433
    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
   434
      |> 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
   435
      |> 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
   436
    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
   437
  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
   438
    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
   439
  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
   440
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
   441
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   442
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
   443
  let
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   444
    fun no_abstr (t $ u) = no_abstr t andalso no_abstr u
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   445
      | no_abstr (Abs (_, _, t)) = no_abstr t
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   446
      | no_abstr (Const (name, _)) = not (Code.is_abstr thy name)
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   447
      | no_abstr _ = true
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   448
    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
   449
      andalso no_abstr (Thm.prop_of eqn)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   450
    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
   451
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   452
  in
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   453
    if is_valid_eq abs_eq_thm then
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   454
      (ABS_EQ, 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
   455
    else
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   456
      let
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   457
        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
   458
      in
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   459
        if rty_body = qty_body then
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   460
          (REP_EQ, 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
   461
        else
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   462
          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
   463
          then
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   464
            (REP_EQ, Code.add_abs_default_eqn (the opt_rep_eq_thm) thy)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   465
          else
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   466
            (NONE_EQ, 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
   467
      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
   468
  end
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   469
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   470
local
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   471
  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
   472
    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
   473
      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
   474
    else
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   475
      if is_Type qty then
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   476
        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
   477
        else 
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   478
          let 
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   479
            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
   480
            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
   481
          in
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   482
            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
   483
          end
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
        true
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   486
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   487
  fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   488
    let
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   489
      fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   490
    in
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   491
      Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty]
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   492
    end
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   493
  
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   494
  exception DECODE
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   495
    
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   496
  fun decode_code_eq thm =
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   497
    if Thm.nprems_of thm > 0 then raise DECODE 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   498
    else
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   499
      let
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   500
        val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   501
        val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   502
        fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   503
      in
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   504
        (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   505
      end
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   506
  
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   507
  structure Data = Generic_Data
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   508
  (
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   509
    type T = code_eq option
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   510
    val empty = NONE
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   511
    val extend = I
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   512
    fun merge _ = NONE
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   513
  );
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   514
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   515
  fun register_encoded_code_eq thm thy =
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   516
    let
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   517
      val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   518
      val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   519
    in
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   520
      Context.theory_map (Data.put (SOME code_eq)) thy
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   521
    end
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   522
    handle DECODE => thy
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   523
  
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   524
  val register_code_eq_attribute = Thm.declaration_attribute
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   525
    (fn thm => Context.mapping (register_encoded_code_eq thm) I)
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   526
  val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute)
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   527
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   528
in
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   529
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   530
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
   531
  let
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   532
    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
   533
  in
60230
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   534
    if no_no_code lthy (rty, qty) then
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   535
      let
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   536
        val lthy = (snd oo Local_Theory.note) 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   537
          ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) lthy
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   538
        val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy))
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   539
        val code_eq = if is_some opt_code_eq then the opt_code_eq 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   540
          else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   541
            which code equation is going to be used. This is going to be resolved at the
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   542
            point when an interpretation of the locale is executed. *)
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   543
        val lthy = Local_Theory.declaration {syntax = false, pervasive = true} 
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   544
          (K (Data.put NONE)) lthy
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   545
      in
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   546
        (code_eq, lthy)
4857d553c52c go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
kuncar
parents: 60229
diff changeset
   547
      end
57663
b590fcd03a4a store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents: 57642
diff changeset
   548
    else
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   549
      (NONE_EQ, lthy)
55724
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   550
  end
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   551
end
7572fc374f80 more robust registration of code equations
kuncar
parents: 55723
diff changeset
   552
            
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   553
(*
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   554
  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
   555
    on a representation type.
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   556
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   557
  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
   558
  qty - an abstract type of the new constant
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   559
  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
   560
  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
   561
    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
   562
  par_thms - a parametricity theorem for rhs
47852
0c3b8d036a5c documentation of the Lifting package on the ML level & tuned
kuncar
parents: 47699
diff changeset
   563
*)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   564
60239
755e11e2e15d make SML/NJ more happy;
wenzelm
parents: 60231
diff changeset
   565
fun add_lift_def (config: config) var qty rhs rsp_thm par_thms lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   566
  let
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   567
    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
   568
    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
   569
    val absrep_trm =  quot_thm_abs quot_thm
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   570
    val rty_forced = (domain_type o fastype_of) absrep_trm
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   571
    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
   572
    val lhs = Free (Binding.name_of (#1 var), qty)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   573
    val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   574
    val (_, prop') = Local_Defs.cert_def lthy prop
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   575
    val (_, newrhs) = Local_Defs.abs_def prop'
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   576
    val var = (#notes config = false ? apfst Binding.concealed) var
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   577
    val def_name = if #notes config then Thm.def_binding (#1 var) else Binding.empty
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   578
    
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   579
    val ((lift_const, (_ , def_thm)), lthy) = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   580
      Local_Theory.define (var, ((def_name, []), newrhs)) lthy
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   581
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   582
    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
   583
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   584
    val abs_eq_thm = generate_abs_eq lthy def_thm rsp_thm quot_thm
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   585
    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
   586
47545
a2850a16e30f Lifting: generate more thms & note them & tuned
kuncar
parents: 47504
diff changeset
   587
    fun qualify defname suffix = Binding.qualified true suffix defname
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   588
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   589
    fun notes names =
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   590
      let
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   591
        val lhs_name = (#1 var)
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   592
        val rsp_thmN = qualify lhs_name "rsp"
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   593
        val abs_eq_thmN = qualify lhs_name "abs_eq"
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   594
        val rep_eq_thmN = qualify lhs_name "rep_eq"
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   595
        val transfer_ruleN = qualify lhs_name "transfer"
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   596
        val notes = 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   597
          [(rsp_thmN, [], [rsp_thm]), 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   598
          (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules),
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   599
          (abs_eq_thmN, [], [abs_eq_thm])] 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   600
          @ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => [])
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   601
      in
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   602
        if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   603
        else map_filter (fn (_, attrs, thms) => if null attrs then NONE 
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   604
          else SOME ((Binding.empty, []), [(thms, attrs)])) notes
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   605
      end
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   606
    val (code_eq, lthy) = register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) lthy
60219
2bce9a690b1e lift_definition: return the result of lifting
kuncar
parents: 60218
diff changeset
   607
    val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm 
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   608
          opt_rep_eq_thm code_eq transfer_rules
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   609
  in
60225
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   610
    lthy
eb4e322734bf note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents: 60224
diff changeset
   611
      |> Local_Theory.notes (notes (#notes config)) |> snd
60224
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   612
      |> ` (fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def)
e759afe46a8c export the result of lifting_def
kuncar
parents: 60221
diff changeset
   613
      ||> Local_Theory.restore
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   614
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   615
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   616
(* 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
   617
  liftings in the current setting *)
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   618
fun get_cr_pcr_eqs ctxt =
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   619
  let
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   620
    fun collect (data : Lifting_Info.quotient) l =
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   621
      if is_some (#pcr_info data) 
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   622
      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
   623
      else l
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   624
    val table = Lifting_Info.get_quotients ctxt
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   625
  in
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   626
    Symtab.fold (fn (_, data) => fn l => collect data l) table []
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   627
  end
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   628
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   629
fun prepare_lift_def add_lift_def var qty rhs par_thms lthy =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   630
  let
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   631
    val rsp_rel = Lifting_Term.equiv_relation lthy (fastype_of rhs, qty)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   632
    val rty_forced = (domain_type o fastype_of) rsp_rel;
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   633
    val forced_rhs = force_rty_type lthy rty_forced rhs;
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   634
    val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv 
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   635
      (Raw_Simplifier.rewrite lthy false (get_cr_pcr_eqs lthy)))
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   636
    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
   637
      |> Thm.cterm_of lthy
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   638
      |> cr_to_pcr_conv
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   639
      |> `Thm.concl_of
56540
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   640
      |>> Logic.dest_equals
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   641
      |>> snd
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   642
    val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2
8267d1ff646f fix the reflexivity prover
kuncar
parents: 56524
diff changeset
   643
    val opt_proven_rsp_thm = try_prove_reflexivity lthy prsp_tm
51994
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   644
    
82cc2aeb7d13 stronger reflexivity prover
kuncar
parents: 51957
diff changeset
   645
    fun after_qed internal_rsp_thm lthy = 
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   646
      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
   647
  in
51374
84d01fd733cf lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents: 51314
diff changeset
   648
    case opt_proven_rsp_thm of
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   649
      SOME thm => (NONE, K (after_qed thm))
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   650
      | NONE => (SOME prsp_tm, after_qed) 
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   651
  end
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
   652
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   653
fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy =
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   654
  let
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   655
    val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   656
  in
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   657
    case goal of
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   658
      SOME goal => 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   659
        let 
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   660
          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
   661
            |> Thm.close_derivation
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   662
        in
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   663
          after_qed rsp_thm lthy
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   664
        end
60231
0daab758e087 don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents: 60230
diff changeset
   665
      | NONE => after_qed Drule.dummy_thm lthy
60218
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   666
  end
7e24e172052e lift_definition: interface also with tactic
kuncar
parents: 60217
diff changeset
   667
60227
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   668
fun lift_def config var qty rhs tac par_thms lthy = gen_lift_def (add_lift_def config)
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   669
  var qty rhs tac par_thms lthy
eacf75e4da95 return also which code equation was used; tuned
kuncar
parents: 60225
diff changeset
   670
53651
ee90c67502c9 restoring Transfer/Lifting context
kuncar
parents: 53207
diff changeset
   671
end (* structure *)