# HG changeset patch # User kuncar # Date 1335442273 -7200 # Node ID 5a10e24fe7abd1de7bd28160f46e1c5e725c2d1f # Parent 7bcdaa255a00d3c267001e0044839d1f76d01a54 tuned; don't generate abs code if quotient_type is used diff -r 7bcdaa255a00 -r 5a10e24fe7ab src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML Thu Apr 26 12:03:11 2012 +0200 +++ b/src/HOL/Tools/Lifting/lifting_setup.ML Thu Apr 26 14:11:13 2012 +0200 @@ -38,8 +38,8 @@ (def_thm, lthy'') end -fun define_abs_type no_abs_code quot_thm lthy = - if not no_abs_code andalso Lifting_Def.can_generate_code_cert quot_thm then +fun define_abs_type gen_abs_code quot_thm lthy = + if gen_abs_code andalso Lifting_Def.can_generate_code_cert quot_thm then let val abs_type_thm = quot_thm RS @{thm Quotient_abs_rep} val add_abstype_attribute = @@ -78,7 +78,7 @@ @ (map Pretty.string_of errs))) end -fun setup_lifting_infr no_abs_code quot_thm lthy = +fun setup_lifting_infr gen_abs_code quot_thm lthy = let val _ = quot_thm_sanity_check lthy quot_thm val (_, qtyp) = Lifting_Term.quot_thm_rty_qty quot_thm @@ -89,10 +89,10 @@ lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) - |> define_abs_type no_abs_code quot_thm + |> define_abs_type gen_abs_code quot_thm end -fun setup_by_quotient no_abs_code quot_thm maybe_reflp_thm lthy = +fun setup_by_quotient gen_abs_code quot_thm maybe_reflp_thm lthy = let val transfer_attr = Attrib.internal (K Transfer.transfer_add) val (_, qty) = Lifting_Term.quot_thm_rty_qty quot_thm @@ -124,10 +124,10 @@ [quot_thm RS @{thm Quotient_right_total}]) |> (snd oo Local_Theory.note) ((qualify "rel_eq_transfer", [transfer_attr]), [quot_thm RS @{thm Quotient_rel_eq_transfer}]) - |> setup_lifting_infr no_abs_code quot_thm + |> setup_lifting_infr gen_abs_code quot_thm end -fun setup_by_typedef_thm no_abs_code typedef_thm lthy = +fun setup_by_typedef_thm gen_abs_code typedef_thm lthy = let val transfer_attr = Attrib.internal (K Transfer.transfer_add) val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o prop_of) typedef_thm @@ -175,10 +175,10 @@ [[quot_thm] MRSL @{thm Quotient_right_unique}]) |> (snd oo Local_Theory.note) ((qualify "right_total", [transfer_attr]), [[quot_thm] MRSL @{thm Quotient_right_total}]) - |> setup_lifting_infr no_abs_code quot_thm + |> setup_lifting_infr gen_abs_code quot_thm end -fun setup_lifting_cmd no_abs_code xthm opt_reflp_xthm lthy = +fun setup_lifting_cmd gen_abs_code xthm opt_reflp_xthm lthy = let val input_thm = singleton (Attrib.eval_thms lthy) xthm val input_term = (HOLogic.dest_Trueprop o prop_of) input_thm @@ -201,14 +201,14 @@ val reflp_thm = singleton (Attrib.eval_thms lthy) reflp_xthm val _ = sanity_check_reflp_thm reflp_thm in - setup_by_quotient no_abs_code input_thm (SOME reflp_thm) lthy + setup_by_quotient gen_abs_code input_thm (SOME reflp_thm) lthy end - | NONE => setup_by_quotient no_abs_code input_thm NONE lthy + | NONE => setup_by_quotient gen_abs_code input_thm NONE lthy fun setup_typedef () = case opt_reflp_xthm of SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." - | NONE => setup_by_typedef_thm no_abs_code input_thm lthy + | NONE => setup_by_typedef_thm gen_abs_code input_thm lthy in case input_term of (Const (@{const_name Quotient}, _) $ _ $ _ $ _ $ _) => setup_quotient () @@ -216,12 +216,12 @@ | _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." end -val opt_no_abs_code = - Scan.optional (@{keyword "("} |-- Parse.!!! ((Parse.reserved "no_abs_code" >> K true) --| @{keyword ")"})) false +val opt_gen_abs_code = + Scan.optional (@{keyword "("} |-- Parse.!!! ((Parse.reserved "no_abs_code" >> K false) --| @{keyword ")"})) true val _ = Outer_Syntax.local_theory @{command_spec "setup_lifting"} "Setup lifting infrastructure" - (opt_no_abs_code -- Parse_Spec.xthm -- Scan.option Parse_Spec.xthm >> - (fn ((no_abs_code, xthm), opt_reflp_xthm) => setup_lifting_cmd no_abs_code xthm opt_reflp_xthm)) + (opt_gen_abs_code -- Parse_Spec.xthm -- Scan.option Parse_Spec.xthm >> + (fn ((gen_abs_code, xthm), opt_reflp_xthm) => setup_lifting_cmd gen_abs_code xthm opt_reflp_xthm)) end;