Nominal/Lift.thy
changeset 2015 3e7969262809
parent 2008 1bddffddc03f
child 2296 45a69c9cc4cc
child 2325 29532d69111c
equal deleted inserted replaced
2014:17684f7eaeb9 2015:3e7969262809
     1 theory Lift
     1 theory Lift
     2 imports "../Nominal-General/Nominal2_Atoms" 
     2 imports "../Nominal-General/Nominal2_Atoms"
     3         "../Nominal-General/Nominal2_Eqvt" 
     3         "../Nominal-General/Nominal2_Eqvt"
     4         "../Nominal-General/Nominal2_Supp" 
     4         "../Nominal-General/Nominal2_Supp"
     5         "Abs" "Perm" "Equivp" "Rsp" "Attic/Fv"
     5         "Abs" "Perm" "Equivp" "Rsp"
     6 begin
     6 begin
     7 
     7 
     8 
     8 
     9 ML {*
     9 ML {*
    10 fun define_quotient_types binds tys alphas equivps ctxt =
    10 fun define_quotient_types binds tys alphas equivps ctxt =
    64 in
    64 in
    65   rename_thm_bvars (un_raw_names (Quotient_Tacs.lifted qtys ctxt thm))
    65   rename_thm_bvars (un_raw_names (Quotient_Tacs.lifted qtys ctxt thm))
    66 end
    66 end
    67 *}
    67 *}
    68 
    68 
    69 
       
    70 
       
    71 
       
    72 ML {*
       
    73 fun define_fv_alpha_export dt binds bns ctxt =
       
    74 let
       
    75   val (((fv_ts_loc, fv_def_loc), ord_fv_ts_loc), ctxt') =
       
    76     define_fv dt binds bns ctxt;
       
    77   val (alpha, ctxt'') =
       
    78     define_alpha dt binds bns fv_ts_loc ctxt';
       
    79   val alpha_ts_loc = #preds alpha
       
    80   val alpha_induct_loc = #induct alpha
       
    81   val alpha_intros_loc = #intrs alpha;
       
    82   val alpha_cases_loc = #elims alpha
       
    83   val morphism = ProofContext.export_morphism ctxt'' ctxt;
       
    84   val fv_ts = map (Morphism.term morphism) fv_ts_loc;
       
    85   val ord_fv_ts = map (Morphism.term morphism) ord_fv_ts_loc;
       
    86   val fv_def = Morphism.fact morphism fv_def_loc;
       
    87   val alpha_ts = map (Morphism.term morphism) alpha_ts_loc;
       
    88   val alpha_induct = Morphism.thm morphism alpha_induct_loc;
       
    89   val alpha_intros = Morphism.fact morphism alpha_intros_loc
       
    90   val alpha_cases = Morphism.fact morphism alpha_cases_loc
       
    91 in
       
    92   ((((fv_ts, ord_fv_ts), fv_def), ((alpha_ts, alpha_intros), (alpha_cases, alpha_induct))), ctxt'')
       
    93 end;
       
    94 *}
       
    95 
       
    96 end
    69 end
    97 
    70