Nominal/Lift.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 18 Mar 2010 07:26:36 +0100
changeset 1494 923413256cbb
parent 1342 2b98012307f7
child 1497 1c9931e5039a
permissions -rw-r--r--
Clean 'Lift', start working only on exported things in Parser.

theory Lift
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp"
begin

ML {*
fun lift_thm ctxt thm =
let
  fun un_raw name = unprefix "_raw" name handle Fail _ => name
  fun add_under names = hd names :: (map (prefix "_") (tl names))
  fun un_raws name = implode (map un_raw (add_under (space_explode "_" name)))
  val un_raw_names = rename_vars un_raws
in
  un_raw_names (snd (Quotient_Tacs.lifted_attrib (Context.Proof ctxt, thm)))
end
*}

ML {*
fun quotient_lift_consts_export spec ctxt =
let
  val (result, ctxt') = fold_map Quotient_Def.quotient_lift_const spec ctxt;
  val (ts_loc, defs_loc) = split_list result;
  val morphism = ProofContext.export_morphism ctxt' ctxt;
  val ts = map (Morphism.term morphism) ts_loc
  val defs = Morphism.fact morphism defs_loc
in
  (ts, defs, ctxt')
end
*}

end