# HG changeset patch # User Christian Urban # Date 1271263640 -7200 # Node ID 2a61e91019a35585b7d7bf380394a7ff7e665928 # Parent fcc660ece04048bd3381981d80e02602715e6560# Parent b435ee87d9c8728504af2c6ad863aeb161c95656 merged diff -r fcc660ece040 -r 2a61e91019a3 Nominal/Fv.thy --- a/Nominal/Fv.thy Wed Apr 14 18:46:59 2010 +0200 +++ b/Nominal/Fv.thy Wed Apr 14 18:47:20 2010 +0200 @@ -416,7 +416,7 @@ *} ML {* -fun define_fv_alpha (dt_info : Datatype_Aux.info) bindsall bns lthy = +fun define_fv (dt_info : Datatype_Aux.info) bindsall bns lthy = let val thy = ProofContext.theory_of lthy; val {descr, sorts, ...} = dt_info; @@ -434,18 +434,8 @@ val (bn_fv_bns, fv_bn_names_eqs) = fv_bns thy dt_info fv_frees rel_bns; val fvbns = map snd bn_fv_bns; val (fv_bn_names, fv_bn_eqs) = split_list fv_bn_names_eqs; - val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) => - "alpha_" ^ name_of_typ (nth_dtyp i)) descr); - val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr; - val alpha_frees = map Free (alpha_names ~~ alpha_types); - (* We assume that a bn is either recursive or not *) - val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns; - val (alpha_bn_names, (bn_alpha_bns, alpha_bn_eqs)) = - alpha_bns dt_info alpha_frees bns bns_rec - val alpha_bn_frees = map snd bn_alpha_bns; - val alpha_bn_types = map fastype_of alpha_bn_frees; - fun fv_alpha_constr ith_dtyp (cname, dts) bindcs = + fun fv_constr ith_dtyp (cname, dts) bindcs = let val Ts = map (typ_of_dtyp descr sorts) dts; val bindslen = length bindcs @@ -457,11 +447,8 @@ val bindcs = map fst bind_pis; val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); val args = map Free (names ~~ Ts); - val names2 = Name.variant_list (pi_strs @ names) (Datatype_Prop.make_tnames Ts); - val args2 = map Free (names2 ~~ Ts); val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); val fv_c = nth fv_frees ith_dtyp; - val alpha = nth alpha_frees ith_dtyp; val arg_nos = 0 upto (length dts - 1) fun fv_bind args (NONE, i, _, _) = if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else @@ -471,7 +458,6 @@ (* TODO goes the code for preiously defined nominal datatypes *) @{term "{} :: atom set"} | fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i) - fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) fun fv_binds_as_set args relevant = mk_union (map (setify o fv_bind args) relevant) fun find_nonrec_binder j (SOME (f, false), i, _, _) = if i = j then SOME f else NONE | find_nonrec_binder _ _ = NONE @@ -498,6 +484,81 @@ end; val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ arg_nos)))) + in + fv_eq + end; + fun fv_eq (i, (_, _, constrs)) binds = map2i (fv_constr i) constrs binds; + val fveqs = map2i fv_eq descr (gather_binds bindsall) + val fv_eqs_perfv = fveqs + val rel_bns_nos = map (fn (_, i, _) => i) rel_bns; + fun filter_fun (_, b) = b mem rel_bns_nos; + val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1)) + val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs))) + val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs))) + val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs); + val fv_names_all = fv_names_fst @ fv_bn_names; + val add_binds = map (fn x => (Attrib.empty_binding, x)) +(* Function_Fun.add_fun Function_Common.default_config ... true *) + val (fvs, lthy') = (Primrec.add_primrec + (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy) + val (fvs2, lthy'') = + if fv_eqs_snd = [] then (([], []), lthy') else + (Primrec.add_primrec + (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy') + val ordered_fvs = fv_frees @ fvbns; + val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) +in + ((all_fvs, ordered_fvs), lthy'') +end +*} + +ML {* +fun define_alpha (dt_info : Datatype_Aux.info) bindsall bns fv_frees lthy = +let + val thy = ProofContext.theory_of lthy; + val {descr, sorts, ...} = dt_info; + fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i); +(* TODO: We need a transitive closure, but instead we do this hack considering + all binding functions as recursive or not *) + val nr_bns = + if (non_rec_binds bindsall) = [] then [] + else map (fn (bn, _, _) => bn) bns; + val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) => + "alpha_" ^ name_of_typ (nth_dtyp i)) descr); + val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr; + val alpha_frees = map Free (alpha_names ~~ alpha_types); + (* We assume that a bn is either recursive or not *) + val bns_rec = map (fn (bn, _, _) => not (bn mem nr_bns)) bns; + val (alpha_bn_names, (bn_alpha_bns, alpha_bn_eqs)) = + alpha_bns dt_info alpha_frees bns bns_rec + val alpha_bn_frees = map snd bn_alpha_bns; + val alpha_bn_types = map fastype_of alpha_bn_frees; + + fun alpha_constr ith_dtyp (cname, dts) bindcs = + let + val Ts = map (typ_of_dtyp descr sorts) dts; + val bindslen = length bindcs + val pi_strs_same = replicate bindslen "pi" + val pi_strs = Name.variant_list [] pi_strs_same; + val pis = map (fn ps => Free (ps, @{typ perm})) pi_strs; + val bind_pis_gath = bindcs ~~ pis; + val bind_pis = un_gather_binds_cons bind_pis_gath; + val names = Name.variant_list pi_strs (Datatype_Prop.make_tnames Ts); + val args = map Free (names ~~ Ts); + val names2 = Name.variant_list (pi_strs @ names) (Datatype_Prop.make_tnames Ts); + val args2 = map Free (names2 ~~ Ts); + val c = Const (cname, Ts ---> (nth_dtyp ith_dtyp)); + val alpha = nth alpha_frees ith_dtyp; + val arg_nos = 0 upto (length dts - 1) + fun fv_bind args (NONE, i, _, _) = + if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else + if ((is_atom thy) o fastype_of) (nth args i) then mk_single_atom (nth args i) else + if ((is_atom_set thy) o fastype_of) (nth args i) then mk_atom_set (nth args i) else + if ((is_atom_fset thy) o fastype_of) (nth args i) then mk_atom_fset (nth args i) else + (* TODO goes the code for preiously defined nominal datatypes *) + @{term "{} :: atom set"} + | fv_bind args (SOME (f, _), i, _, _) = f $ (nth args i) + fun fv_binds args relevant = mk_union (map (fv_bind args) relevant) val alpha_rhs = HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2))); fun alpha_arg ((dt, arg_no), (arg, arg2)) = @@ -572,36 +633,20 @@ fold (fn pi_str => fn t => HOLogic.mk_exists (pi_str, @{typ perm}, t)) pi_strs alpha_lhss val alpha_eq = Logic.mk_implies (HOLogic.mk_Trueprop alpha_lhss_ex, alpha_rhs) in - (fv_eq, alpha_eq) + alpha_eq end; - fun fv_alpha_eq (i, (_, _, constrs)) binds = map2i (fv_alpha_constr i) constrs binds; - val fveqs_alphaeqs = map2i fv_alpha_eq descr (gather_binds bindsall) - val (fv_eqs_perfv, alpha_eqs) = apsnd flat (split_list (map split_list fveqs_alphaeqs)) - val rel_bns_nos = map (fn (_, i, _) => i) rel_bns; - fun filter_fun (_, b) = b mem rel_bns_nos; - val all_fvs = (fv_names ~~ fv_eqs_perfv) ~~ (0 upto (length fv_names - 1)) - val (fv_names_fst, fv_eqs_fst) = apsnd flat (split_list (map fst (filter_out filter_fun all_fvs))) - val (fv_names_snd, fv_eqs_snd) = apsnd flat (split_list (map fst (filter filter_fun all_fvs))) - val fv_eqs_all = fv_eqs_fst @ (flat fv_bn_eqs); - val fv_names_all = fv_names_fst @ fv_bn_names; + fun alpha_eq (i, (_, _, constrs)) binds = map2i (alpha_constr i) constrs binds; + val alphaeqs = map2i alpha_eq descr (gather_binds bindsall) + val alpha_eqs = flat alphaeqs val add_binds = map (fn x => (Attrib.empty_binding, x)) -(* Function_Fun.add_fun Function_Common.default_config ... true *) - val (fvs, lthy') = (Primrec.add_primrec - (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_all) (add_binds fv_eqs_all) lthy) - val (fvs2, lthy'') = - if fv_eqs_snd = [] then (([], []), lthy') else - (Primrec.add_primrec - (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names_snd) (add_binds fv_eqs_snd) lthy') - val (alphas, lthy''') = (Inductive.add_inductive_i + val (alphas, lthy') = (Inductive.add_inductive_i {quiet_mode = true, verbose = false, alt_name = Binding.empty, coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false} (map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) (alpha_names @ alpha_bn_names) (alpha_types @ alpha_bn_types)) [] - (add_binds (alpha_eqs @ flat alpha_bn_eqs)) [] lthy'') - val ordered_fvs = fv_frees @ fvbns; - val all_fvs = (fst fvs @ fst fvs2, snd fvs @ snd fvs2) + (add_binds (alpha_eqs @ flat alpha_bn_eqs)) [] lthy) in - (((all_fvs, ordered_fvs), alphas), lthy''') + (alphas, lthy') end *} diff -r fcc660ece040 -r 2a61e91019a3 Nominal/Lift.thy --- a/Nominal/Lift.thy Wed Apr 14 18:46:59 2010 +0200 +++ b/Nominal/Lift.thy Wed Apr 14 18:47:20 2010 +0200 @@ -1,7 +1,7 @@ theory Lift imports "../Nominal-General/Nominal2_Atoms" "../Nominal-General/Nominal2_Eqvt" - "../Nominal_General/Nominal2_Supp" + "../Nominal-General/Nominal2_Supp" "Abs" "Perm" "Equivp" "Rsp" begin @@ -69,13 +69,16 @@ ML {* fun define_fv_alpha_export dt binds bns ctxt = let - val ((((fv_ts_loc, fv_def_loc), ord_fv_ts_loc), alpha), ctxt') = - define_fv_alpha dt binds bns ctxt; + val (((fv_ts_loc, fv_def_loc), ord_fv_ts_loc), ctxt') = + define_fv dt binds bns ctxt; + val fv_ts_nobn = take (length bns) fv_ts_loc + val (alpha, ctxt'') = + define_alpha dt binds bns fv_ts_nobn ctxt'; val alpha_ts_loc = #preds alpha val alpha_induct_loc = #induct alpha val alpha_intros_loc = #intrs alpha; val alpha_cases_loc = #elims alpha - val morphism = ProofContext.export_morphism ctxt' ctxt; + val morphism = ProofContext.export_morphism ctxt'' ctxt; val fv_ts = map (Morphism.term morphism) fv_ts_loc; val ord_fv_ts = map (Morphism.term morphism) ord_fv_ts_loc; val fv_def = Morphism.fact morphism fv_def_loc; @@ -84,7 +87,7 @@ val alpha_intros = Morphism.fact morphism alpha_intros_loc val alpha_cases = Morphism.fact morphism alpha_cases_loc in - ((((fv_ts, ord_fv_ts), fv_def), ((alpha_ts, alpha_intros), (alpha_cases, alpha_induct))), ctxt') + ((((fv_ts, ord_fv_ts), fv_def), ((alpha_ts, alpha_intros), (alpha_cases, alpha_induct))), ctxt'') end; *}