arch_split: added optional definition override for crunch. Reduced qualification commands to minimal required set.

This commit is contained in:
Daniel Matichuk 2016-05-03 14:23:30 +10:00
parent ec399ad38e
commit 670d1c118d
7 changed files with 271 additions and 407 deletions

View File

@ -13,6 +13,8 @@ imports "wp/NonDetMonadVCG"
keywords "crunch" "crunch_ignore" :: thy_decl
begin
named_theorems "crunch_def"
ML_file "crunch-cmd.ML"
ML_file "Crunch.ML"

View File

@ -10,7 +10,7 @@
theory Crunch_Test
imports Crunch Crunch_Test_Qualified GenericLib
imports Crunch Crunch_Test_Qualified GenericLib Defs
begin
text {* Test cases for crunch *}
@ -137,6 +137,21 @@ definition
crunch test[wp]: crunch_foo9 "\<lambda>x. x > y" (ignore: modify)
definition
"crunch_foo10 (x :: nat) \<equiv> do
modify (op + x);
modify (op + x)
od"
(*crunch_def attribute overrides definition lookup *)
lemma crunch_foo10_def2[crunch_def]:
"crunch_foo10 = crunch_foo9"
unfolding crunch_foo10_def[abs_def] crunch_foo9_def[abs_def]
by simp
crunch test[wp]: crunch_foo10 "\<lambda>x. x > y"
(* crunch_ignore works within a locale *)
crunch_ignore (add: modify)

View File

@ -13,45 +13,17 @@ keywords "qualify" :: thy_decl and "end_qualify" :: thy_decl
begin
ML \<open>
type qualify_args = {name : string, deep : bool, target_name : string}
type qualify_args = {name : string, target_name : string}
structure Data = Theory_Data
(
type T = (theory * qualify_args) option *
((string *
((binding Symtab.table) * (* facts *)
(binding Symtab.table) * (* consts *)
(binding Symtab.table) (* types *))) list);
val empty = (NONE, []);
type T = (theory * qualify_args) option;
val empty = NONE;
val extend = I;
fun merge (((_, tabs), (_, tabs')) : T * T) =
(NONE, AList.join (op =)
(fn _ => fn ((facts, consts, types), (facts', consts', types')) =>
(Symtab.merge (op =) (facts, facts'),
Symtab.merge (op =) (consts, consts'),
Symtab.merge (op =) (types, types')))
(tabs, tabs'));
fun merge ((_, _) : T * T) = NONE;
);
fun get_qualify thy = fst (Data.get thy);
fun get_tabs_of thy str =
the_default (Symtab.empty, Symtab.empty, Symtab.empty) (AList.lookup (op =) (snd (Data.get thy)) str);
fun get_facts_of thy str = #1 (get_tabs_of thy str);
fun get_consts_of thy str = #2 (get_tabs_of thy str);
fun get_types_of thy str = #3 (get_tabs_of thy str);
fun map_tabs_of str f thy =
Data.map (apsnd (AList.map_default (op =) (str, (Symtab.empty, Symtab.empty, Symtab.empty)) f)) thy;
fun map_facts_of str f thy = map_tabs_of str (@{apply 3(1)} f) thy;
fun map_consts_of str f thy = map_tabs_of str (@{apply 3(2)} f) thy;
fun map_types_of str f thy = map_tabs_of str (@{apply 3(3)} f) thy;
fun get_qualify thy = Data.get thy;
fun make_bind space thy nm =
let
@ -119,48 +91,18 @@ fun make_bind_local nm =
fun set_global_qualify (args : qualify_args) thy =
let
val str = #name args
val _ = Locale.check thy (#target_name args, Position.none)
val _ = case get_qualify thy of SOME _ => error "Already in a qualify block!" | NONE => ();
val thy' = Data.map (apfst (K (SOME (thy,args)))) thy;
val thy' = Data.map (K (SOME (thy,args))) thy;
in if #deep args then
let
val lthy = Named_Target.begin (#target_name args, Position.none) thy';
val facts =
Facts.fold_static (fn (nm, _) => add_qualified str nm) (Global_Theory.facts_of thy) []
|> map (`make_bind_local)
val const_space = #const_space (Consts.dest (Proof_Context.consts_of lthy));
val consts = fold (fn (nm, _) => add_qualified str nm) (#constants (Consts.dest (Sign.consts_of thy))) []
|> map (`(make_bind const_space thy'))
val types =
Name_Space.fold_table (fn (nm, _) => add_qualified str nm) (#types (Type.rep_tsig (Sign.tsig_of thy))) []
|> map (`make_bind_local)
val thy'' = thy'
|> map_facts_of str (fold (fn (b, nm) => (Symtab.update (nm, b))) facts)
|> map_consts_of str (fold (fn (b, nm) => (Symtab.update (nm, b))) consts)
|> map_types_of str (fold (fn (b, nm) => (Symtab.update (nm, b))) types)
val thy''' = thy''
|> fold_rev (fn (nm, b) => Global_Theory.alias_fact b nm) (Symtab.dest (get_facts_of thy'' str))
|> fold_rev (fn (nm, b) => Sign.const_alias b nm) (Symtab.dest (get_consts_of thy'' str))
|> fold_rev (fn (nm, b) => Sign.type_alias b nm) (Symtab.dest (get_types_of thy'' str))
in thy''' end
else thy'
end
in thy' end
val _ =
Outer_Syntax.command @{command_keyword qualify} "begin global qualification"
(Parse.name -- (Parse.opt_target -- Args.mode "deep")>>
(fn (str, (target, deep)) =>
Toplevel.theory (set_global_qualify {name = str, deep = deep, target_name = case target of SOME (nm, _) => nm | _ => str})));
(Parse.name -- Parse.opt_target>>
(fn (str, target) =>
Toplevel.theory (set_global_qualify {name = str, target_name = case target of SOME (nm, _) => nm | _ => str})));
fun syntax_alias global_alias local_alias b name =
Local_Theory.declaration {syntax = true, pervasive = true} (fn phi =>
@ -188,20 +130,10 @@ fun end_global_qualify thy =
val types = get_new_types old_thy (Sign.tsig_of thy);
val thy' = thy
|> map_facts_of nm (fold (fn (b, nm) => (Symtab.update (nm, b))) facts)
|> map_consts_of nm (fold (fn (b, nm) => (Symtab.update (nm, b))) consts)
|> map_types_of nm (fold (fn (b, nm) => (Symtab.update (nm, b))) types);
val facts' = (if #deep args then (map swap (Symtab.dest (get_facts_of thy' nm))) else facts);
val consts' = (if #deep args then (map swap (Symtab.dest (get_consts_of thy' nm))) else consts);
val types' = (if #deep args then (map swap (Symtab.dest (get_types_of thy' nm))) else types);
val thy'' = thy'
|> (fn thy => fold (Global_Theory.hide_fact false o snd) facts' thy)
|> (fn thy => fold (Sign.hide_const false o snd) consts' thy)
|> (fn thy => fold (Sign.hide_type false o snd) types' thy);
val thy'' = thy
|> (fn thy => fold (Global_Theory.hide_fact false o snd) facts thy)
|> (fn thy => fold (Sign.hide_const false o snd) consts thy)
|> (fn thy => fold (Sign.hide_type false o snd) types thy);
val lthy = Named_Target.begin (#target_name args, Position.none) thy''
|> Local_Theory.map_background_naming (Name_Space.parent_path #> Name_Space.mandatory_path nm);
@ -211,7 +143,7 @@ fun end_global_qualify thy =
|> fold (uncurry const_alias) consts
|> fold (uncurry type_alias) types;
in Local_Theory.exit_global lthy' |> Data.map (apfst (K NONE)) end
in Local_Theory.exit_global lthy' |> Data.map (K NONE) end
val _ =
Outer_Syntax.command @{command_keyword end_qualify} "end global qualification"

219
lib/Requalify.thy Normal file
View File

@ -0,0 +1,219 @@
(*
Requalify constants, types and facts into the current naming
*)
theory Requalify
imports Main
keywords "requalify_facts" :: thy_decl and "requalify_types" :: thy_decl and "requalify_consts" :: thy_decl and
"global_naming" :: thy_decl
begin
ML \<open>
local
fun all_facts_of ctxt =
let
val thy = Proof_Context.theory_of ctxt;
val global_facts = Global_Theory.facts_of thy;
in
Facts.dest_static false [] global_facts
end;
fun global_fact ctxt nm =
let
val facts = Proof_Context.facts_of ctxt;
val {name, thms, ...} = (Facts.retrieve (Context.Proof ctxt) facts (nm, Position.none));
fun tl' (_ :: xs) = xs
| tl' _ = []
fun matches suf (gnm, gthms) =
let
val gsuf = Long_Name.explode gnm |> tl' |> tl' |> Long_Name.implode;
in suf = gsuf andalso eq_list Thm.eq_thm_prop (thms, gthms)
end
in
case Long_Name.dest_local name of NONE => (name, thms) | SOME suf =>
(case (find_first (matches suf) (all_facts_of ctxt)) of
SOME x => x
| NONE => raise Fail ("Couldn't find global equivalent of local fact: " ^ nm))
end
fun syntax_alias global_alias local_alias b (name : string) =
Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
let val b' = Morphism.binding phi b
in Context.mapping (global_alias b' name) (local_alias b' name) end);
val fact_alias = syntax_alias Global_Theory.alias_fact Proof_Context.fact_alias;
val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias;
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;
in
fun gen_requalify get_proper_nm parse_nm check_nm alias =
(Parse.opt_target -- Scan.repeat1 (Parse.position (Scan.ahead parse_nm -- Parse.xname))
>> (fn (target,bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun read_entry ((entry, t), pos) lthy =
let
val local_nm = get_proper_nm lthy t;
val _ = check_nm lthy (entry, (local_nm, pos));
val b = Binding.make (Long_Name.base_name t, pos)
val lthy' = lthy
|> alias b local_nm
in lthy' end
in fold read_entry bs lthy end)))
local
val get_const_nm = ((fst o dest_Const) oo (Proof_Context.read_const {proper = true, strict = false}))
val get_type_nm = ((fst o dest_Type) oo (Proof_Context.read_type_name {proper = true, strict = false}))
val get_fact_nm = (fst oo global_fact)
fun check_fact lthy (_, (nm, pos)) = Proof_Context.get_fact lthy (Facts.Named ((nm,pos), NONE))
in
val _ =
Outer_Syntax.command @{command_keyword requalify_consts} "alias const with current naming"
(gen_requalify get_const_nm Parse.const (fn lthy => fn (e, _) => get_const_nm lthy e) const_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_types} "alias type with current naming"
(gen_requalify get_type_nm Parse.typ (fn lthy => fn (e, _) => get_type_nm lthy e) type_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_facts} "alias fact with current naming"
(gen_requalify get_fact_nm Parse.xthm check_fact fact_alias)
val _ =
Outer_Syntax.command @{command_keyword global_naming} "change global naming of context block"
(Parse.name >> (fn naming =>
Toplevel.local_theory NONE NONE
(Local_Theory.map_background_naming (Name_Space.parent_path #> Name_Space.mandatory_path naming))))
end
end
\<close>
(*Tests and examples *)
locale Requalify_Locale
begin
typedecl requalify_type
definition "requalify_const == (undefined :: requalify_type)"
end
typedecl requalify_type
definition "requalify_const == (undefined :: requalify_type)"
context Requalify_Locale begin global_naming Requalify_Locale2
requalify_consts requalify_const
requalify_types requalify_type
requalify_facts requalify_const_def
end
term "requalify_const :: requalify_type"
term "Requalify_Locale2.requalify_const :: Requalify_Locale2.requalify_type"
lemma "Requalify_Locale2.requalify_const = (undefined :: Requalify_Locale2.requalify_type)"
by (simp add: Requalify_Locale2.requalify_const_def)
consts requalify_test_f :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const Requalify.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
context Requalify_Locale begin
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const Requalify.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f2)?
apply (rule f1)
apply (simp add: requalify_const_def)
done
end
context Requalify_Locale begin global_naming global
requalify_consts Requalify.requalify_const
requalify_types Requalify.requalify_type
requalify_facts Requalify.requalify_const_def
end
context Requalify_Locale begin
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const global.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
end
context begin interpretation Requalify_Locale .
lemma
assumes f1: "requalify_test_f requalify_const Requalify_Locale2.requalify_const"
and f2: "requalify_test_f Requalify_Locale2.requalify_const global.requalify_const"
shows "requalify_test_f Requalify_Locale2.requalify_const requalify_const" "requalify_const = undefined"
apply (rule f1)?
apply (rule f2)
apply (simp add: requalify_const_def)
done
end
locale Requalify_Locale3
begin
typedecl requalify_type2
definition "requalify_const2 == (undefined :: requalify_type2)"
end
context begin interpretation Requalify_Locale3 .
requalify_consts requalify_const2
requalify_types requalify_type2
requalify_facts requalify_const2_def
end
lemma "(requalify_const2 :: requalify_type2) = undefined"
by (simp add: requalify_const2_def)
context Requalify_Locale3 begin
lemma "(requalify_const2 :: requalify_type2) = undefined"
by (simp add: requalify_const2_def)
end
end

View File

@ -1,321 +0,0 @@
(*
Drop mandatory qualifiers from internal locale facts, constants and types
Optionally add a different qualifier
*)
theory Unqualify
imports Main
keywords "unqualify_facts" :: thy_decl and "unqualify_consts" :: thy_decl and "unqualify_types" :: thy_decl and
"shadow_facts" :: thy_decl and "shadow_types" :: thy_decl and "shadow_consts" :: thy_decl and
"requalify_facts" :: thy_decl and "requalify_types" :: thy_decl and "requalify_consts" :: thy_decl and
"global_naming" :: thy_decl
begin
ML \<open>
local
fun map_option _ NONE = NONE
| map_option f (SOME x) = SOME (f x)
fun make_binding_raw (nm, pos) =
let
val path = Long_Name.explode nm |> rev;
val b' = fold (Binding.qualify true) (tl path) (Binding.make (hd path, pos));
in b' end
fun make_binding b =
let
val nm = Binding.name_of b;
val pos = Binding.pos_of b;
in make_binding_raw (nm, pos) end
fun make_notation b =
let
val str = Binding.name_of b
|> String.translate(fn #"_" => "'_" | x => Char.toString x)
in str end
fun all_facts_of ctxt =
let
val thy = Proof_Context.theory_of ctxt;
val global_facts = Global_Theory.facts_of thy;
in
Facts.dest_static false [] global_facts
end;
fun global_fact ctxt nm =
let
val facts = Proof_Context.facts_of ctxt;
val {name, thms, ...} = (Facts.retrieve (Context.Proof ctxt) facts (nm, Position.none));
fun tl' (_ :: xs) = xs
| tl' _ = []
fun matches suf (gnm, gthms) =
let
val gsuf = Long_Name.explode gnm |> tl' |> tl' |> Long_Name.implode;
in suf = gsuf andalso eq_list Thm.eq_thm_prop (thms, gthms)
end
in
case Long_Name.dest_local name of NONE => (name, thms) | SOME suf =>
(case (find_first (matches suf) (all_facts_of ctxt)) of
SOME x => x
| NONE => raise Fail ("Couldn't find global equivalent of local fact: " ^ nm))
end
fun syntax_alias global_alias local_alias b (name : string) =
Local_Theory.declaration {syntax = false, pervasive = true} (fn phi =>
let val b' = Morphism.binding phi b
in Context.mapping (global_alias b' name) (local_alias b' name) end);
val fact_alias = syntax_alias Global_Theory.alias_fact Proof_Context.fact_alias;
val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias;
val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;
in
val _ =
Outer_Syntax.command @{command_keyword unqualify_facts} "unqualify facts"
(Parse.opt_target -- Scan.option (Args.parens Parse.name) -- Parse.xthms1 >>
(fn ((target,qual),rfs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
val facts = Proof_Context.facts_of lthy;
val nms =
map (fn (rf, atts) => (Facts.select rf []; ((Facts.name_of_ref rf, Facts.pos_of_ref rf), atts))) rfs;
fun retrieve ((nm, pos), _) = Facts.retrieve (Context.Proof lthy) facts (nm, pos);
val thmss = map (`retrieve) nms
|> map (fn ({thms, static, ...}, b) =>
(if not static then error ("Can't unqualify dynamic fact: " ^ (fst (fst b)) ^ Position.here (snd (fst b)))
else (apfst make_binding_raw b, thms)));
val naming = Sign.naming_of (Proof_Context.theory_of lthy)
|> the_default I (map_option Name_Space.mandatory_path qual);
fun check_attribs ctxt raw_srcs = let
val srcs = map (Attrib.check_src ctxt) raw_srcs;
val _ = map (Attrib.attribute ctxt) srcs;
in srcs end;
val lthy' = Local_Theory.background_theory (fn thy =>
let
val thy' = Context.theory_map (Name_Space.map_naming (K naming)) thy;
val global_ctxt = (Proof_Context.init_global thy);
val morph = Local_Theory.standard_morphism lthy global_ctxt;
val thmss' = map (fn ((b, att),thms) => ((b, []), [(thms, check_attribs global_ctxt att)])) thmss;
in
snd (Attrib.global_notes Thm.theoremK (Attrib.transform_facts morph thmss') thy')
end) lthy
in
lthy'
end)))
val _ =
Outer_Syntax.command @{command_keyword unqualify_consts} "unqualify consts"
(Parse.opt_target -- Scan.option (Args.parens Parse.name) --
Scan.repeat1 (Parse.position (Parse.const -- Scan.option (Parse.$$$ "::" |-- Parse.typ))) >> (fn ((target,qual),bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun read_const (t, T_in) =
let
val (nm, T) = dest_Const (Proof_Context.read_const {proper = true, strict = false} lthy t);
in case map_option (Syntax.read_typ lthy) T_in of
SOME T' => (nm, T')
| NONE => (nm, T) end
val ts = map (`( read_const o fst)) bs;
fun get_const (((nm, T) , (_, pos))) = ((nm, pos), Const (nm, T))
val ts' = map (apfst make_binding_raw o get_const) ts;
val naming = Sign.naming_of (Proof_Context.theory_of lthy)
|> the_default I (map_option Name_Space.mandatory_path qual);
val lthy' = Local_Theory.background_theory (fn thy =>
let
val thy' = Context.theory_map (Name_Space.map_naming (K naming)) thy;
val export = Morphism.term (Local_Theory.standard_morphism lthy (Proof_Context.init_global thy));
in
fold (fn (b, t) =>
let
val b' = b
|> make_binding;
val t' = export t;
in
Sign.add_abbrev Print_Mode.internal(b', t') #> snd #>
Sign.notation true ("", false) [(t', Delimfix (make_notation b'))]
end) ts' thy'
end) lthy
in
lthy'
end)))
val _ =
Outer_Syntax.command @{command_keyword unqualify_types} "unqualify types"
(Parse.opt_target -- Scan.option (Args.parens Parse.name) --
Scan.repeat1 (Parse.position Parse.type_const) >> (fn ((target,qual),bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun err (T, pos) =
error ("Not a defined type or type synonym: " ^ Syntax.string_of_typ lthy T ^ Position.here pos)
val Ts = map (`(Proof_Context.read_type_name {proper = true, strict = false} lthy o fst)) bs;
fun get_type (T, (_, pos)) = case try dest_Type T of
SOME (nm, Ts) => (if can dest_funT T then err (T,pos) else ((nm, pos), (nm,Ts)))
| NONE => err (T,pos)
val Ts' = map (apfst make_binding_raw o get_type) Ts;
val naming = Sign.naming_of (Proof_Context.theory_of lthy)
|> the_default I (map_option Name_Space.mandatory_path qual);
val lthy' = Local_Theory.background_theory (fn thy =>
let
val thy' = Context.theory_map (Name_Space.map_naming (K naming)) thy;
in
fold (fn (b, (nm,frees_raw)) =>
let
val Ts = lthy
|> Variable.invent_types (map (fn _ => Proof_Context.default_sort lthy ("'a",~1)) frees_raw)
|> fst
|> map TFree
val T = Type (nm, Ts)
val frees = map (fst o dest_TFree) Ts
val b' = make_binding b;
val str = make_notation b'
|> fold (fn _ => fn s => "_ " ^ s) Ts
in
Sign.add_type_abbrev lthy (b',frees, T) #>
Sign.type_notation true ("", false) [(T, Delimfix str)]
end) Ts' thy'
end) lthy
in
lthy'
end)))
fun gen_shadow get_proper_nm alias =
(Parse.opt_target -- Scan.repeat1 (Parse.position (Parse.name))
>> (fn (target,bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun read_entry (t, pos) lthy =
let
val global_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy);
val (global_nm : string) = get_proper_nm global_ctxt t;
val local_nm = get_proper_nm lthy t;
val _ = if local_nm = global_nm then error (global_nm ^ " is already globally defined." ^ Position.here pos) else ()
val b = Binding.make (t, pos)
val global_naming = Proof_Context.naming_of global_ctxt;
val lthy' = lthy
|> Local_Theory.map_background_naming (K global_naming)
|> alias b global_nm
|> Local_Theory.restore_background_naming lthy
in lthy' end
in fold read_entry bs lthy end)))
fun gen_requalify get_proper_nm alias =
(Parse.opt_target -- Scan.repeat1 (Parse.position (Parse.name))
>> (fn (target,bs) =>
Toplevel.local_theory NONE target (fn lthy =>
let
fun read_entry (t, pos) lthy =
let
val local_nm = get_proper_nm lthy t;
val global_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy);
val opt_global_nm = try (get_proper_nm global_ctxt) t;
val _ = if SOME local_nm = opt_global_nm
then error (local_nm ^ " is already globally accessible." ^ Position.here pos)
else ()
val b = Binding.make (t, pos)
val lthy' = lthy
|> alias b local_nm
in lthy' end
in fold read_entry bs lthy end)))
local
val get_const_nm = ((fst o dest_Const) oo (Proof_Context.read_const {proper = true, strict = false}))
val get_type_nm = ((fst o dest_Type) oo (Proof_Context.read_type_name {proper = true, strict = false}))
val get_fact_nm = (fst oo global_fact)
in
val _ =
Outer_Syntax.command @{command_keyword shadow_consts} "shadow local consts with global ones"
(gen_shadow get_const_nm const_alias)
val _ =
Outer_Syntax.command @{command_keyword shadow_types} "shadow local types with global ones"
(gen_shadow get_type_nm type_alias)
val _ =
Outer_Syntax.command @{command_keyword shadow_facts} "shadow local facts with global ones"
(gen_shadow get_fact_nm fact_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_consts} "alias local const with current naming"
(gen_requalify get_const_nm const_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_types} "alias local type with current naming"
(gen_requalify get_type_nm type_alias)
val _ =
Outer_Syntax.command @{command_keyword requalify_facts} "alias local fact with current naming"
(gen_requalify get_fact_nm fact_alias)
val _ =
Outer_Syntax.command @{command_keyword global_naming} "change global naming of context block"
(Parse.name >> (fn naming =>
Toplevel.local_theory NONE NONE
(Local_Theory.map_background_naming (Name_Space.parent_path #> Name_Space.mandatory_path naming))))
end
end
\<close>
end

View File

@ -205,6 +205,23 @@ val unfold_get_params = @{thms Let_def return_bind returnOk_bindE
K_bind_def split_def bind_assoc bindE_assoc
trans[OF liftE_bindE return_bind]};
fun def_from_ctxt ctxt const =
let
val crunch_defs = Named_Theorems.get ctxt "Crunch.crunch_def"
val abs_def = Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def;
fun do_filter thm =
let
val (Const (nm, _), _) = Logic.dest_equals (Thm.prop_of (abs_def thm));
in nm = const end
in
case crunch_defs |> filter do_filter of
[x] => [x]
| [] => []
| _ => raise Fail ("Multiple definitions declared for:" ^ const)
end
fun unfold ctxt const triple nmspce =
let
val _ = debug_trace "unfold"
@ -419,7 +436,7 @@ fun crunch cfg pre extra stack const' thms =
let
val ctxt' = Variable.auto_fixes goal ctxt;
val cgoal = goal |> var_precond |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt';
val unfolds' = filter (fn (a, _) => a = const) (#unfolds cfg)
val unfolds' = filter (fn (a, _) => a = const) (#unfolds cfg) @ (map (pair "") (def_from_ctxt ctxt const))
val (ms, unfold_rule) = unfold_data ctxt' const cgoal (#nmspce cfg) unfolds'
|>> map (fn t => (real_const_from_name (fst (dest_Const (head_of t))) (#nmspce cfg) ctxt', t))
|>> subtract (fn (a, b) => a = (fst b))

View File

@ -10,7 +10,7 @@
theory Setup_Locale
imports "../../../lib/Qualify" "../../../lib/Unqualify"
imports "../../../lib/Qualify" "../../../lib/Requalify"
begin
(*