diff -r 51f75d24bd73 -r 5e95387bef45 Nominal/Nominal2.thy --- a/Nominal/Nominal2.thy Thu Jan 06 14:53:38 2011 +0000 +++ b/Nominal/Nominal2.thy Thu Jan 06 19:57:57 2011 +0000 @@ -163,27 +163,14 @@ end *} -ML {* -(* for testing porposes - to exit the procedure early *) -exception TEST of Proof.context - -val (STEPS, STEPS_setup) = Attrib.config_int "STEPS" (K 100); - -fun get_STEPS ctxt = Config.get ctxt STEPS -*} - - -setup STEPS_setup ML {* fun nominal_datatype2 opt_thms_name dts bn_funs bn_eqs bclauses lthy = let (* definition of the raw datatypes *) - val _ = warning "Definition of raw datatypes"; + val _ = trace_msg (K "Defining raw datatypes...") val (raw_dt_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy0) = - if get_STEPS lthy > 0 - then define_raw_dts dts bn_funs bn_eqs bclauses lthy - else raise TEST lthy + define_raw_dts dts bn_funs bn_eqs bclauses lthy val dtinfo = Datatype.the_info (ProofContext.theory_of lthy0) (hd raw_dt_names) val {descr, sorts, ...} = dtinfo @@ -209,186 +196,125 @@ |> `(fn thms => (length thms) div 2) |> uncurry drop - (* definitions of raw permutations by primitive recursion *) - val _ = warning "Definition of raw permutations"; + val _ = trace_msg (K "Defining raw permutations...") val ((raw_perm_funs, raw_perm_simps, raw_perm_laws), lthy2a) = - if get_STEPS lthy0 > 0 - then define_raw_perms raw_full_ty_names raw_tys tvs (flat raw_constrs) raw_induct_thm lthy0 - else raise TEST lthy0 + define_raw_perms raw_full_ty_names raw_tys tvs (flat raw_constrs) raw_induct_thm lthy0 (* noting the raw permutations as eqvt theorems *) val (_, lthy3) = Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_perm_simps) lthy2a - (* definition of raw fv and bn functions *) - val _ = warning "Definition of raw fv- and bn-functions"; + + val _ = trace_msg (K "Defining raw fv- and bn-functions...") val (raw_bns, raw_bn_defs, raw_bn_info, raw_bn_induct, lthy3a) = - if get_STEPS lthy3 > 1 - then define_raw_bns raw_full_ty_names raw_dts raw_bn_funs raw_bn_eqs + define_raw_bns raw_full_ty_names raw_dts raw_bn_funs raw_bn_eqs (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3 - else raise TEST lthy3 - + (* defining the permute_bn functions *) val (raw_perm_bns, raw_perm_bn_simps, lthy3b) = - if get_STEPS lthy3a > 2 - then define_raw_bn_perms raw_tys raw_bn_info raw_cns_info + define_raw_bn_perms raw_tys raw_bn_info raw_cns_info (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3a - else raise TEST lthy3a - + val (raw_fvs, raw_fv_bns, raw_fv_defs, raw_fv_bns_induct, lthy3c) = - if get_STEPS lthy3b > 3 - then define_raw_fvs raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses + define_raw_fvs raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses (raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3b - else raise TEST lthy3b - - (* definition of raw alphas *) - val _ = warning "Definition of alphas"; + + val _ = trace_msg (K "Defining alpha relations...") val (alpha_trms, alpha_bn_trms, alpha_intros, alpha_cases, alpha_induct, lthy4) = - if get_STEPS lthy3c > 4 - then define_raw_alpha raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses raw_fvs lthy3c - else raise TEST lthy3c + define_raw_alpha raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses raw_fvs lthy3c + val alpha_tys = map (domain_type o fastype_of) alpha_trms - (* definition of alpha-distinct lemmas *) - val _ = warning "Distinct theorems"; + val _ = trace_msg (K "Proving distinct theorems...") val alpha_distincts = mk_alpha_distincts lthy4 alpha_cases raw_distinct_thms alpha_trms raw_tys - (* definition of alpha_eq_iff lemmas *) - val _ = warning "Eq-iff theorems"; + val _ = trace_msg (K "Proving eq-iff theorems...") val alpha_eq_iff = - if get_STEPS lthy > 5 - then mk_alpha_eq_iff lthy4 alpha_intros raw_distinct_thms raw_inject_thms alpha_cases - else raise TEST lthy4 - - (* proving equivariance lemmas for bns, fvs, size and alpha *) - val _ = warning "Proving equivariance"; + mk_alpha_eq_iff lthy4 alpha_intros raw_distinct_thms raw_inject_thms alpha_cases + + val _ = trace_msg (K "Proving equivariance of bns, fvs, size and alpha...") val raw_bn_eqvt = - if get_STEPS lthy > 6 - then raw_prove_eqvt raw_bns raw_bn_induct (raw_bn_defs @ raw_perm_simps) lthy4 - else raise TEST lthy4 - + raw_prove_eqvt raw_bns raw_bn_induct (raw_bn_defs @ raw_perm_simps) lthy4 + (* noting the raw_bn_eqvt lemmas in a temprorary theory *) val lthy_tmp = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_bn_eqvt) lthy4) val raw_fv_eqvt = - if get_STEPS lthy > 7 - then raw_prove_eqvt (raw_fvs @ raw_fv_bns) raw_fv_bns_induct (raw_fv_defs @ raw_perm_simps) + raw_prove_eqvt (raw_fvs @ raw_fv_bns) raw_fv_bns_induct (raw_fv_defs @ raw_perm_simps) (Local_Theory.restore lthy_tmp) - else raise TEST lthy4 - + val raw_size_eqvt = - if get_STEPS lthy > 8 - then raw_prove_eqvt raw_size_trms raw_induct_thms (raw_size_thms @ raw_perm_simps) + raw_prove_eqvt raw_size_trms raw_induct_thms (raw_size_thms @ raw_perm_simps) (Local_Theory.restore lthy_tmp) |> map (rewrite_rule @{thms permute_nat_def[THEN eq_reflection]}) |> map (fn thm => thm RS @{thm sym}) - else raise TEST lthy4 - + val lthy5 = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_fv_eqvt) lthy_tmp) val (alpha_eqvt, lthy6) = - if get_STEPS lthy > 9 - then Nominal_Eqvt.equivariance true (alpha_trms @ alpha_bn_trms) alpha_induct alpha_intros lthy5 - else raise TEST lthy4 + Nominal_Eqvt.equivariance true (alpha_trms @ alpha_bn_trms) alpha_induct alpha_intros lthy5 - (* proving alpha equivalence *) - val _ = warning "Proving equivalence" - + val _ = trace_msg (K "Proving equivalence of alpha...") val alpha_refl_thms = - if get_STEPS lthy > 10 - then raw_prove_refl alpha_trms alpha_bn_trms alpha_intros raw_induct_thm lthy6 - else raise TEST lthy6 + raw_prove_refl alpha_trms alpha_bn_trms alpha_intros raw_induct_thm lthy6 val alpha_sym_thms = - if get_STEPS lthy > 11 - then raw_prove_sym (alpha_trms @ alpha_bn_trms) alpha_intros alpha_induct lthy6 - else raise TEST lthy6 + raw_prove_sym (alpha_trms @ alpha_bn_trms) alpha_intros alpha_induct lthy6 val alpha_trans_thms = - if get_STEPS lthy > 12 - then raw_prove_trans (alpha_trms @ alpha_bn_trms) (raw_distinct_thms @ raw_inject_thms) - alpha_intros alpha_induct alpha_cases lthy6 - else raise TEST lthy6 + raw_prove_trans (alpha_trms @ alpha_bn_trms) (raw_distinct_thms @ raw_inject_thms) + alpha_intros alpha_induct alpha_cases lthy6 val (alpha_equivp_thms, alpha_bn_equivp_thms) = - if get_STEPS lthy > 13 - then raw_prove_equivp alpha_trms alpha_bn_trms alpha_refl_thms alpha_sym_thms - alpha_trans_thms lthy6 - else raise TEST lthy6 + raw_prove_equivp alpha_trms alpha_bn_trms alpha_refl_thms alpha_sym_thms + alpha_trans_thms lthy6 - (* proving alpha implies alpha_bn *) - val _ = warning "Proving alpha implies bn" - + val _ = trace_msg (K "Proving alpha implies bn...") val alpha_bn_imp_thms = - if get_STEPS lthy > 14 - then raw_prove_bn_imp alpha_trms alpha_bn_trms alpha_intros alpha_induct lthy6 - else raise TEST lthy6 + raw_prove_bn_imp alpha_trms alpha_bn_trms alpha_intros alpha_induct lthy6 - (* respectfulness proofs *) + val _ = trace_msg (K "Proving respectfulness...") val raw_funs_rsp_aux = - if get_STEPS lthy > 15 - then raw_fv_bn_rsp_aux alpha_trms alpha_bn_trms raw_fvs + raw_fv_bn_rsp_aux alpha_trms alpha_bn_trms raw_fvs raw_bns raw_fv_bns alpha_induct (raw_bn_defs @ raw_fv_defs) lthy6 - else raise TEST lthy6 - val raw_funs_rsp = - if get_STEPS lthy > 16 - then map mk_funs_rsp raw_funs_rsp_aux - else raise TEST lthy6 + val raw_funs_rsp = map mk_funs_rsp raw_funs_rsp_aux val raw_size_rsp = - if get_STEPS lthy > 17 - then - raw_size_rsp_aux (alpha_trms @ alpha_bn_trms) alpha_induct + raw_size_rsp_aux (alpha_trms @ alpha_bn_trms) alpha_induct (raw_size_thms @ raw_size_eqvt) lthy6 |> map mk_funs_rsp - else raise TEST lthy6 val raw_constrs_rsp = - if get_STEPS lthy > 18 - then raw_constrs_rsp (flat raw_constrs) alpha_trms alpha_intros + raw_constrs_rsp (flat raw_constrs) alpha_trms alpha_intros (alpha_bn_imp_thms @ raw_funs_rsp_aux) lthy6 - else raise TEST lthy6 - - val alpha_permute_rsp = - if get_STEPS lthy > 19 - then map mk_alpha_permute_rsp alpha_eqvt - else raise TEST lthy6 + + val alpha_permute_rsp = map mk_alpha_permute_rsp alpha_eqvt val alpha_bn_rsp = - if get_STEPS lthy > 20 - then raw_alpha_bn_rsp alpha_bn_trms alpha_bn_equivp_thms alpha_bn_imp_thms - else raise TEST lthy6 + raw_alpha_bn_rsp alpha_bn_trms alpha_bn_equivp_thms alpha_bn_imp_thms val raw_perm_bn_rsp = - if get_STEPS lthy > 21 - then raw_perm_bn_rsp (alpha_trms @ alpha_bn_trms) raw_perm_bns alpha_induct + raw_perm_bn_rsp (alpha_trms @ alpha_bn_trms) raw_perm_bns alpha_induct alpha_intros raw_perm_bn_simps lthy6 - else raise TEST lthy6 (* noting the quot_respects lemmas *) val (_, lthy6a) = - if get_STEPS lthy > 22 - then Local_Theory.note ((Binding.empty, [rsp_attr]), + Local_Theory.note ((Binding.empty, [rsp_attr]), raw_constrs_rsp @ raw_funs_rsp @ raw_size_rsp @ alpha_permute_rsp @ alpha_bn_rsp @ raw_perm_bn_rsp) lthy6 - else raise TEST lthy6 - (* defining the quotient type *) - val _ = warning "Declaring the quotient types" + val _ = trace_msg (K "Defining the quotient types...") val qty_descr = map (fn (vs, bind, mx, _) => (vs, bind, mx)) dts val (qty_infos, lthy7) = - if get_STEPS lthy > 23 - then define_qtypes qty_descr alpha_tys alpha_trms alpha_equivp_thms lthy6a - else raise TEST lthy6a + define_qtypes qty_descr alpha_tys alpha_trms alpha_equivp_thms lthy6a val qtys = map #qtyp qty_infos val qty_full_names = map (fst o dest_Type) qtys val qty_names = map Long_Name.base_name qty_full_names - (* defining of quotient term-constructors, binding functions, free vars functions *) - val _ = warning "Defining the quotient constants" + val _ = trace_msg (K "Defining the quotient constants...") val qconstrs_descrs = (map2 o map2) (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) (get_cnstrs dts) raw_constrs @@ -415,8 +341,6 @@ val ((((((qconstrs_infos, qbns_info), qfvs_info), qfv_bns_info), qalpha_bns_info), qperm_bns_info), lthy8) = - if get_STEPS lthy > 24 - then lthy7 |> fold_map (define_qconsts qtys) qconstrs_descrs ||>> define_qconsts qtys qbns_descr @@ -424,18 +348,12 @@ ||>> define_qconsts qtys qfv_bns_descr ||>> define_qconsts qtys qalpha_bns_descr ||>> define_qconsts qtys qperm_bn_descr - else raise TEST lthy7 - (* definition of the quotient permfunctions and pt-class *) val lthy9 = - if get_STEPS lthy > 25 - then define_qperms qtys qty_full_names tvs qperm_descr raw_perm_laws lthy8 - else raise TEST lthy8 + define_qperms qtys qty_full_names tvs qperm_descr raw_perm_laws lthy8 val lthy9a = - if get_STEPS lthy > 26 - then define_qsizes qtys qty_full_names tvs qsize_descr lthy9 - else raise TEST lthy9 + define_qsizes qtys qty_full_names tvs qsize_descr lthy9 val qtrms = (map o map) #qconst qconstrs_infos val qbns = map #qconst qbns_info @@ -444,64 +362,45 @@ val qalpha_bns = map #qconst qalpha_bns_info val qperm_bns = map #qconst qperm_bns_info - (* lifting of the theorems *) - val _ = warning "Lifting of Theorems" + val _ = trace_msg (K "Lifting of theorems...") val eq_iff_simps = @{thms alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel_def prod.cases} val ((((((qdistincts, qeq_iffs), qfv_defs), qbn_defs), qperm_simps), qfv_qbn_eqvts), lthyA) = - if get_STEPS lthy > 27 - then - lthy9a - |> lift_thms qtys [] alpha_distincts - ||>> lift_thms qtys eq_iff_simps alpha_eq_iff - ||>> lift_thms qtys [] raw_fv_defs - ||>> lift_thms qtys [] raw_bn_defs - ||>> lift_thms qtys [] raw_perm_simps - ||>> lift_thms qtys [] (raw_fv_eqvt @ raw_bn_eqvt) - else raise TEST lthy9a + lthy9a + |> lift_thms qtys [] alpha_distincts + ||>> lift_thms qtys eq_iff_simps alpha_eq_iff + ||>> lift_thms qtys [] raw_fv_defs + ||>> lift_thms qtys [] raw_bn_defs + ||>> lift_thms qtys [] raw_perm_simps + ||>> lift_thms qtys [] (raw_fv_eqvt @ raw_bn_eqvt) val ((((((qsize_eqvt, [qinduct]), qexhausts), qsize_simps), qperm_bn_simps), qalpha_refl_thms), lthyB) = - if get_STEPS lthy > 28 - then - lthyA - |> lift_thms qtys [] raw_size_eqvt - ||>> lift_thms qtys [] [raw_induct_thm] - ||>> lift_thms qtys [] raw_exhaust_thms - ||>> lift_thms qtys [] raw_size_thms - ||>> lift_thms qtys [] raw_perm_bn_simps - ||>> lift_thms qtys [] alpha_refl_thms - else raise TEST lthyA + lthyA + |> lift_thms qtys [] raw_size_eqvt + ||>> lift_thms qtys [] [raw_induct_thm] + ||>> lift_thms qtys [] raw_exhaust_thms + ||>> lift_thms qtys [] raw_size_thms + ||>> lift_thms qtys [] raw_perm_bn_simps + ||>> lift_thms qtys [] alpha_refl_thms val qinducts = Project_Rule.projections lthyA qinduct - (* supports lemmas *) - val _ = warning "Proving Supports Lemmas and fs-Instances" + val _ = trace_msg (K "Proving supp lemmas and fs-instances...") val qsupports_thms = - if get_STEPS lthy > 29 - then prove_supports lthyB qperm_simps (flat qtrms) - else raise TEST lthyB + prove_supports lthyB qperm_simps (flat qtrms) (* finite supp lemmas *) - val qfsupp_thms = - if get_STEPS lthy > 30 - then prove_fsupp lthyB qtys qinduct qsupports_thms - else raise TEST lthyB + val qfsupp_thms = prove_fsupp lthyB qtys qinduct qsupports_thms (* fs instances *) - val lthyC = - if get_STEPS lthy > 31 - then fs_instance qtys qty_full_names tvs qfsupp_thms lthyB - else raise TEST lthyB + val lthyC = fs_instance qtys qty_full_names tvs qfsupp_thms lthyB - (* fv - supp equality *) - val _ = warning "Proving Equality between fv and supp" + val _ = trace_msg (K "Proving equality between fv and supp...") val qfv_supp_thms = - if get_STEPS lthy > 32 - then prove_fv_supp qtys (flat qtrms) qfvs qfv_bns qalpha_bns qfv_defs qeq_iffs + prove_fv_supp qtys (flat qtrms) qfvs qfv_bns qalpha_bns qfv_defs qeq_iffs qperm_simps qfv_qbn_eqvts qinduct (flat raw_bclauses) lthyC - else [] (* postprocessing of eq and fv theorems *) @@ -524,31 +423,24 @@ |> map (simplify (HOL_basic_ss addsimps transform_thms)) (* proving that the qbn result is finite *) - val qbn_finite_thms = - if get_STEPS lthy > 33 - then prove_bns_finite qtys qbns qinduct qbn_defs lthyC - else [] + val qbn_finite_thms = prove_bns_finite qtys qbns qinduct qbn_defs lthyC (* proving that perm_bns preserve alpha *) val qperm_bn_alpha_thms = - if get_STEPS lthy > 33 - then prove_perm_bn_alpha_thms qtys qperm_bns qalpha_bns qinduct qperm_bn_simps qeq_iffs' + prove_perm_bn_alpha_thms qtys qperm_bns qalpha_bns qinduct qperm_bn_simps qeq_iffs' qalpha_refl_thms lthyC - else [] (* proving the relationship of bn and permute_bn *) val qpermute_bn_thms = - if get_STEPS lthy > 33 - then prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qfv_qbn_eqvts lthyC - else [] + prove_permute_bn_thms qtys qbns qperm_bns qinduct qperm_bn_simps qbn_defs qfv_qbn_eqvts lthyC - (* proving the strong exhausts and induct lemmas *) + val _ = trace_msg (K "Proving strong exhaust lemmas...") val qstrong_exhaust_thms = prove_strong_exhausts lthyC qexhausts bclauses qbn_finite_thms qeq_iffs' qfv_qbn_eqvts qpermute_bn_thms qperm_bn_alpha_thms + val _ = trace_msg (K "Proving strong induct lemmas...") val qstrong_induct_thms = prove_strong_induct lthyC qinduct qstrong_exhaust_thms qsize_simps bclauses - (* noting the theorems *) (* generating the prefix for the theorem names *) @@ -582,7 +474,7 @@ ||>> Local_Theory.note ((thms_suffix "permute_bn", []), qpermute_bn_thms) in lthy9' -end handle TEST ctxt => ctxt +end *} @@ -808,6 +700,10 @@ (main_parser >> nominal_datatype2_cmd) *} +ML {* +trace := true +*} + end