diff -r ff887850d83c -r 92dc6cfa3a95 Nominal/NewParser.thy --- a/Nominal/NewParser.thy Wed Aug 25 20:19:10 2010 +0800 +++ b/Nominal/NewParser.thy Wed Aug 25 22:55:42 2010 +0800 @@ -471,24 +471,50 @@ val qfvs = map #qconst qfvs_info val qfv_bns = map #qconst qfv_bns_info val qalpha_bns = map #qconst qalpha_bns_info + + (* lifting of the theorems *) + val _ = warning "Lifting of Theorems" + + val eq_iff_simps = @{thms alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel.simps + prod.cases} + + val ((((((qdistincts, qeq_iffs), qfv_defs), qbn_defs), qperm_simps), qfv_qbn_eqvts), lthyA) = + if get_STEPS lthy > 20 + 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 + + val (((qsize_eqvt, [qinduct]), qexhausts), lthyB) = + if get_STEPS lthy > 20 + then + lthyA + |> lift_thms qtys [] raw_size_eqvt + ||>> lift_thms qtys [] [raw_induct_thm] + ||>> lift_thms qtys [] raw_exhaust_thms + else raise TEST lthyA + (* temporary theorem bindings *) - val (_, lthy9') = lthy9a - |> Local_Theory.note ((@{binding "distinct"}, []), alpha_distincts) - ||>> Local_Theory.note ((@{binding "eq_iff"}, []), alpha_eq_iff) - ||>> Local_Theory.note ((@{binding "eq_iff_simps"}, []), alpha_eq_iff_simps) - ||>> Local_Theory.note ((@{binding "fv_defs"}, []), raw_fv_defs) - ||>> Local_Theory.note ((@{binding "bn_defs"}, []), raw_bn_defs) - ||>> Local_Theory.note ((@{binding "perm_simps"}, []), raw_perm_simps) - ||>> Local_Theory.note ((@{binding "perm_laws"}, []), raw_perm_laws) - ||>> Local_Theory.note ((@{binding "alpha_bn_imps"}, []), alpha_bn_imp_thms) - ||>> Local_Theory.note ((@{binding "alpha_equivp"}, []), alpha_equivp_thms) - ||>> Local_Theory.note ((@{binding "fv_eqvt"}, []), raw_fv_eqvt) - ||>> Local_Theory.note ((@{binding "bn_eqvt"}, []), raw_bn_eqvt) - ||>> Local_Theory.note ((@{binding "size_eqvt"}, []), raw_size_eqvt) + val (_, lthy9') = lthyB + |> Local_Theory.note ((@{binding "distinct"}, []), qdistincts) + ||>> Local_Theory.note ((@{binding "eq_iff"}, []), qeq_iffs) + ||>> Local_Theory.note ((@{binding "fv_defs"}, []), qfv_defs) + ||>> Local_Theory.note ((@{binding "bn_defs"}, []), qbn_defs) + ||>> Local_Theory.note ((@{binding "perm_simps"}, []), qperm_simps) + ||>> Local_Theory.note ((@{binding "fv_bn_eqvt"}, []), qfv_qbn_eqvts) + ||>> Local_Theory.note ((@{binding "size_eqvt"}, []), qsize_eqvt) + ||>> Local_Theory.note ((@{binding "induct"}, []), [qinduct]) + ||>> Local_Theory.note ((@{binding "exhaust"}, []), qexhausts) + val _ = - if get_STEPS lthy > 20 + if get_STEPS lthy > 21 then true else raise TEST lthy9' (* old stuff *) @@ -553,7 +579,7 @@ fun suffix_bind s = Binding.qualify true q_name (Binding.name s); val _ = warning "Lifting induction"; val constr_names = map (Long_Name.base_name o fst o dest_Const) []; - val q_induct = Rule_Cases.name constr_names (lift_thm lthy13 qtys [] raw_induct_thm); + val q_induct = Rule_Cases.name constr_names (the_single (fst (lift_thms qtys [] [raw_induct_thm] lthy13))); fun note_suffix s th ctxt = snd (Local_Theory.note ((suffix_bind s, []), th) ctxt); fun note_simp_suffix s th ctxt = @@ -563,24 +589,24 @@ [Rule_Cases.name constr_names q_induct]) lthy13; val q_inducts = Project_Rule.projects lthy13 (1 upto (length raw_fvs)) q_induct val (_, lthy14a) = Local_Theory.note ((suffix_bind "inducts", []), q_inducts) lthy14; - val q_perm = map (lift_thm lthy14 qtys []) raw_perm_simps; + val q_perm = fst (lift_thms qtys [] raw_perm_simps lthy14); val lthy15 = note_simp_suffix "perm" q_perm lthy14a; - val q_fv = map (lift_thm lthy15 qtys []) raw_fv_defs; + val q_fv = fst (lift_thms qtys [] raw_fv_defs lthy15); val lthy16 = note_simp_suffix "fv" q_fv lthy15; - val q_bn = map (lift_thm lthy16 qtys []) raw_bn_defs; + val q_bn = fst (lift_thms qtys [] raw_bn_defs lthy16); val lthy17 = note_simp_suffix "bn" q_bn lthy16; val _ = warning "Lifting eq-iff"; (*val _ = map tracing (map PolyML.makestring alpha_eq_iff);*) val eq_iff_unfolded0 = map (Local_Defs.unfold lthy17 @{thms alphas}) alpha_eq_iff val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms Pair_eqvt}) eq_iff_unfolded0 - val q_eq_iff_pre0 = map (lift_thm lthy17 qtys []) eq_iff_unfolded1; + val q_eq_iff_pre0 = fst (lift_thms qtys [] eq_iff_unfolded1 lthy17); val q_eq_iff_pre1 = map (Local_Defs.fold lthy17 @{thms Pair_eqvt}) q_eq_iff_pre0 val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alphas}) q_eq_iff_pre1 val q_eq_iff = map (Local_Defs.unfold lthy17 (Quotient_Info.id_simps_get lthy17)) q_eq_iff_pre2 val (_, lthy18) = Local_Theory.note ((suffix_bind "eq_iff", []), q_eq_iff) lthy17; - val q_dis = map (lift_thm lthy18 qtys []) alpha_distincts; + val q_dis = fst (lift_thms qtys [] alpha_distincts lthy18); val lthy19 = note_simp_suffix "distinct" q_dis lthy18; - val q_eqvt = map (lift_thm lthy19 qtys []) (raw_bn_eqvt @ raw_fv_eqvt); + val q_eqvt = fst (lift_thms qtys [] (raw_bn_eqvt @ raw_fv_eqvt) lthy19); val (_, lthy20) = Local_Theory.note ((Binding.empty, [Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add)]), q_eqvt) lthy19; val _ = warning "Supports";