Nominal/Lift.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 26 Feb 2010 10:34:04 +0100
changeset 1276 3365fce80f0f
parent 1274 d867021d8ac1
child 1277 6eacf60ce41d
permissions -rw-r--r--
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Lift
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
atom_decl ident
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
datatype rtrm2 =
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
  rVr2 "name"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| rAp2 "rtrm2" "rtrm2"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
| rLt2 "ras" "rtrm2" --"bind (bv2 l) in (r)"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
and ras =
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  rAs "name" "rtrm2"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
primrec rbv2 where "rbv2 (rAs x t) = {atom x}"
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
ML {*
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
fun build_eqvts_ funs perms simps induct ctxt =
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
let
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
  val pi = Free ("pi", @{typ perm});
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
  val types = map (domain_type o fastype_of) funs;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
  val fv_indnames = Datatype_Prop.make_tnames (map body_type types);
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
  val args = map Free (fv_indnames ~~ types);
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
  val perm_at = Const (@{const_name permute}, @{typ "perm \<Rightarrow> atom set \<Rightarrow> atom set"})
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  fun eqvtc (fv, (arg, perm)) =
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
    HOLogic.mk_eq ((perm_at $ pi $ (fv $ arg)), (fv $ (perm $ pi $ arg)))
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
  val gl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map eqvtc (funs ~~ (args ~~ perms))))
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
  fun tac _ = (indtac induct fv_indnames THEN_ALL_NEW
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
    (asm_full_simp_tac (HOL_ss addsimps 
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
      (@{thm atom_eqvt} :: (Nominal_ThmDecls.get_eqvts_thms ctxt) @ simps)))) 1
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
in
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
  Goal.prove ctxt ("pi" :: fv_indnames) [] gl tac
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
end
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
*}
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
ML {*
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
print_depth 500;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
val thy1 = @{theory};
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
val tnames = ["rtrm2", "ras"];
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
val ftname = "Lift.rtrm2"
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    43
val ntnames = [(@{binding trm2}, @{typ rtrm2}), (@{binding as}, @{typ ras})]
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
val binds = [[[[]], [[], []], [[], [(SOME @{term rbv2}, 0)]]], 
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
             [[[], []]]  (*, [[], [[], []]] *) ];
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
val bv_simps = @{thms rbv2.simps}
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
val info = Datatype.the_info @{theory} ftname;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
val descr = #descr info;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
val all_full_tnames = map (fn (_, (n, _, _)) => n) descr;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
val full_tnames = List.take (all_full_tnames, length tnames);
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
val induct = #induct info;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
val inducts = #inducts info;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
val infos = map (Datatype.the_info @{theory}) all_full_tnames;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
val inject = flat (map #inject infos);
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
val distinct = flat (map #distinct infos);
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
val ((raw_perm_def, raw_perm_simps, perms), thy2) = define_raw_perms tnames full_tnames thy1;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
val lthy1 = Theory_Target.init NONE thy2
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
val (((fv_ts_loc, fv_def_loc), alpha), lthy2) = define_fv_alpha ftname binds lthy1;
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    59
val alpha_ts_loc = #preds alpha
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
val alpha_intros = #intrs alpha
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
val alpha_cases = #elims alpha
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
val alpha_induct = #induct alpha
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
val alpha_inj = build_alpha_inj alpha_intros (inject @ distinct) alpha_cases lthy2
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
val fv_def = ProofContext.export lthy2 lthy1 fv_def_loc
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
val morphism = ProofContext.export_morphism lthy2 lthy1
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
val fv_ts = map (Morphism.term morphism) fv_ts_loc
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    67
val alpha_ts = map (Morphism.term morphism) alpha_ts_loc
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
val (bv_eqvts, lthy3) = build_eqvts @{binding bv_eqvts} [@{term rbv2}] [hd (tl perms)] 
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  (bv_simps @ raw_perm_def) @{thm rtrm2_ras.inducts(2)} lthy2;
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
val (fv_eqvts, lthy4) = build_eqvts @{binding fv_eqvts} fv_ts_loc perms (fv_def_loc @ raw_perm_def) induct lthy3;
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    71
val alpha_eqvt = build_alpha_eqvts alpha_ts_loc perms (raw_perm_def @ alpha_inj) alpha_induct lthy4;
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    72
val alpha_equivp_loc = build_equivps alpha_ts_loc induct alpha_induct inject alpha_inj distinct alpha_cases alpha_eqvt lthy4;
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    73
val alpha_equivp = ProofContext.export lthy4 lthy1 alpha_equivp_loc
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    74
val lthy5 = define_quotient_type
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    75
  (map (fn ((b, t), alpha) => (([], b, NoSyn), (t, alpha))) (ntnames ~~ alpha_ts))
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    76
  (ALLGOALS (resolve_tac alpha_equivp)) lthy4;
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    77
val ((c, def), lthy6) = Quotient_Def.quotient_lift_const ("Vr2", @{term rVr2}) lthy5
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    78
*}
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    80
ML {*
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    81
val thyf = Local_Theory.exit_global lthy6;
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
*}
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    83
setup {* fn _ => thyf *}
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    84
print_theorems
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    85
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
ML {*
1276
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    90
val lthy5 = define_quotient_type
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    91
  (map (fn ((b, t), alpha) => (([], b, NoSyn), (t, alpha))) (ntnames ~~ alpha_ts))
3365fce80f0f To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    92
  (ALLGOALS (resolve_tac alpha_equivp)) lthy4
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
*}
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
ML {*
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
*}
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99