theory Examples
imports Complex_Main Lattice_Syntax
begin
section {* Fundamental Isabelle types
and statically embedded logical entities *}
text {* Example 1 *}
lemmas latticify = inf_set_eq [symmetric] sup_set_eq [symmetric]
(*bot_set_eq [symmetric]*) (*top_set_eq [symmetric]*)
Inf_set_eq [symmetric] Sup_set_eq [symmetric]
ML {*
val latticify = MetaSimplifier.rewrite_rule
(map Simpdata.mk_eq @{thms latticify})
*}
thm Diff_Int_distrib
ML {*
latticify @{thm Diff_Int_distrib}
*}
text {* Example 2 *}
ML {*
@{term "{ f x | x. P x }"}
*}
ML {*
case @{term "{ f x | x. P x }"}
of _ $ Abs (_, _, t) => t
*}
section {* Passing states and accumulating results *}
text {* Example 3 *}
(*local_setup {* fn lthy =>
snd (LocalTheory.define Thm.definitionK
((@{binding identity}, NoSyn),
((@{binding indentity_def}, []), @{term "\<lambda>x\<Colon>'a. x"})) lthy)
*}*)
print_theorems
text {* Example 4 *}
local_setup {* fn lthy => lthy
|> LocalTheory.define Thm.definitionK
((@{binding identity}, NoSyn),
((@{binding indentity_def}, []), @{term "\<lambda>x\<Colon>'a. x"}))
|> snd
*}
text {* Example 5 *}
term "(default, distinct)"
setup {* fn thy =>
thy
|> fold (Sign.add_const_constraint)
[(@{const_name default}, SOME @{typ "'a::type"}),
(@{const_name distinct}, SOME @{typ "'a::eq list \<Rightarrow> bool"})]
*}
term "(default, distinct)"
text {* Example 6 *}
local_setup {* fn lthy =>
let
fun mk_dummy_var v = ("", TFree (v, @{sort type}));
fun mk_proj i = ("proj_" ^ string_of_int i,
list_abs (map mk_dummy_var (Name.invent_list [] "'a" i), Bound 0));
val projs = map mk_proj (1 upto 9);
in
lthy
|> fold_map (fn (bname, t) => LocalTheory.define Thm.definitionK
((Binding.name bname, NoSyn),
((Binding.name (Thm.def_name bname), []), t)))
projs
|-> (fn defs => LocalTheory.note Thm.lemmaK
((@{binding proj_fold}, []), map (symmetric o snd o snd) defs))
|> snd
end
*}
print_theorems
section {* Managing generic data *}
ML {*
signature REWRITER =
sig
val add: thm -> Context.generic -> Context.generic
val del: thm -> Context.generic -> Context.generic
val rewrite: Context.generic -> conv
end
*}
ML {*
structure Rewriter : REWRITER =
struct
structure Data = GenericDataFun(
struct
type T = thm list;
val empty = [];
fun merge _ = Thm.merge_thms;
val extend = I;
end);
fun add thm = Data.map (Thm.add_thm thm);
fun del thm = Data.map (Thm.del_thm thm);
fun rewrite ctxt = MetaSimplifier.rewrite false
(map Simpdata.mk_eq (Data.get ctxt));
end
*}
attribute_setup rewrite =
{* Scan.succeed (Thm.declaration_attribute Rewriter.add) *}
""
method_setup rewrite =
{* Scan.state >> (fn ctxt =>
K (SIMPLE_METHOD' (CONVERSION (Rewriter.rewrite ctxt)))) *}
""
lemma "A \<sqinter> B = \<Inter>{A, B}"
proof -
note latticify [symmetric, rewrite]
show ?thesis
apply rewrite
apply simp
done
qed
lemma "A \<inter> B = A \<sqinter> B"
proof -
show ?thesis
apply rewrite
oops
section {* For your amusement *}
end