|
1 theory Examples |
|
2 imports Complex_Main Lattice_Syntax |
|
3 begin |
|
4 |
|
5 section {* Fundamental Isabelle types |
|
6 and statically embedded logical entities *} |
|
7 |
|
8 text {* Example 1 *} |
|
9 |
|
10 lemmas latticify = inf_set_eq [symmetric] sup_set_eq [symmetric] |
|
11 (*bot_set_eq [symmetric]*) (*top_set_eq [symmetric]*) |
|
12 Inf_set_eq [symmetric] Sup_set_eq [symmetric] |
|
13 |
|
14 ML {* |
|
15 val latticify = MetaSimplifier.rewrite_rule |
|
16 (map Simpdata.mk_eq @{thms latticify}) |
|
17 *} |
|
18 |
|
19 thm Diff_Int_distrib |
|
20 |
|
21 ML {* |
|
22 latticify @{thm Diff_Int_distrib} |
|
23 *} |
|
24 |
|
25 text {* Example 2 *} |
|
26 |
|
27 ML {* |
|
28 @{term "{ f x | x. P x }"} |
|
29 *} |
|
30 |
|
31 ML {* |
|
32 case @{term "{ f x | x. P x }"} |
|
33 of _ $ Abs (_, _, t) => t |
|
34 *} |
|
35 |
|
36 |
|
37 section {* Passing states and accumulating results *} |
|
38 |
|
39 text {* Example 3 *} |
|
40 |
|
41 (*local_setup {* fn lthy => |
|
42 snd (LocalTheory.define Thm.definitionK |
|
43 ((@{binding identity}, NoSyn), |
|
44 ((@{binding indentity_def}, []), @{term "\<lambda>x\<Colon>'a. x"})) lthy) |
|
45 *}*) |
|
46 |
|
47 print_theorems |
|
48 |
|
49 text {* Example 4 *} |
|
50 |
|
51 local_setup {* fn lthy => lthy |
|
52 |> LocalTheory.define Thm.definitionK |
|
53 ((@{binding identity}, NoSyn), |
|
54 ((@{binding indentity_def}, []), @{term "\<lambda>x\<Colon>'a. x"})) |
|
55 |> snd |
|
56 *} |
|
57 |
|
58 text {* Example 5 *} |
|
59 |
|
60 term "(default, distinct)" |
|
61 |
|
62 setup {* fn thy => |
|
63 thy |
|
64 |> fold (Sign.add_const_constraint) |
|
65 [(@{const_name default}, SOME @{typ "'a::type"}), |
|
66 (@{const_name distinct}, SOME @{typ "'a::eq list \<Rightarrow> bool"})] |
|
67 *} |
|
68 |
|
69 term "(default, distinct)" |
|
70 |
|
71 text {* Example 6 *} |
|
72 |
|
73 local_setup {* fn lthy => |
|
74 let |
|
75 fun mk_dummy_var v = ("", TFree (v, @{sort type})); |
|
76 fun mk_proj i = ("proj_" ^ string_of_int i, |
|
77 list_abs (map mk_dummy_var (Name.invent_list [] "'a" i), Bound 0)); |
|
78 val projs = map mk_proj (1 upto 9); |
|
79 in |
|
80 lthy |
|
81 |> fold_map (fn (bname, t) => LocalTheory.define Thm.definitionK |
|
82 ((Binding.name bname, NoSyn), |
|
83 ((Binding.name (Thm.def_name bname), []), t))) |
|
84 projs |
|
85 |-> (fn defs => LocalTheory.note Thm.lemmaK |
|
86 ((@{binding proj_fold}, []), map (symmetric o snd o snd) defs)) |
|
87 |> snd |
|
88 end |
|
89 *} |
|
90 |
|
91 print_theorems |
|
92 |
|
93 |
|
94 section {* Managing generic data *} |
|
95 |
|
96 ML {* |
|
97 signature REWRITER = |
|
98 sig |
|
99 val add: thm -> Context.generic -> Context.generic |
|
100 val del: thm -> Context.generic -> Context.generic |
|
101 val rewrite: Context.generic -> conv |
|
102 end |
|
103 *} |
|
104 |
|
105 ML {* |
|
106 structure Rewriter : REWRITER = |
|
107 struct |
|
108 |
|
109 structure Data = GenericDataFun( |
|
110 struct |
|
111 type T = thm list; |
|
112 val empty = []; |
|
113 fun merge _ = Thm.merge_thms; |
|
114 val extend = I; |
|
115 end); |
|
116 |
|
117 fun add thm = Data.map (Thm.add_thm thm); |
|
118 fun del thm = Data.map (Thm.del_thm thm); |
|
119 |
|
120 fun rewrite ctxt = MetaSimplifier.rewrite false |
|
121 (map Simpdata.mk_eq (Data.get ctxt)); |
|
122 |
|
123 end |
|
124 *} |
|
125 |
|
126 attribute_setup rewrite = |
|
127 {* Scan.succeed (Thm.declaration_attribute Rewriter.add) *} |
|
128 "" |
|
129 |
|
130 method_setup rewrite = |
|
131 {* Scan.state >> (fn ctxt => |
|
132 K (SIMPLE_METHOD' (CONVERSION (Rewriter.rewrite ctxt)))) *} |
|
133 "" |
|
134 |
|
135 lemma "A \<sqinter> B = \<Inter>{A, B}" |
|
136 proof - |
|
137 note latticify [symmetric, rewrite] |
|
138 show ?thesis |
|
139 apply rewrite |
|
140 apply simp |
|
141 done |
|
142 qed |
|
143 |
|
144 lemma "A \<inter> B = A \<sqinter> B" |
|
145 proof - |
|
146 show ?thesis |
|
147 apply rewrite |
|
148 oops |
|
149 |
|
150 section {* For your amusement *} |
|
151 |
|
152 end |