author | wenzelm |
Wed, 28 Aug 2024 22:54:45 +0200 | |
changeset 80786 | 70076ba563d2 |
parent 69597 | ff784d5a5bfb |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
13208 | 1 |
(* Title: HOL/Prolog/Test.thy |
2 |
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) |
|
3 |
*) |
|
9015 | 4 |
|
63167 | 5 |
section \<open>Basic examples\<close> |
9015 | 6 |
|
17311 | 7 |
theory Test |
8 |
imports HOHH |
|
9 |
begin |
|
9015 | 10 |
|
17311 | 11 |
typedecl nat |
12 |
||
13 |
typedecl 'a list |
|
9015 | 14 |
|
17311 | 15 |
consts |
16 |
Nil :: "'a list" ("[]") |
|
17 |
Cons :: "'a => 'a list => 'a list" (infixr "#" 65) |
|
9015 | 18 |
|
80786
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
19 |
text \<open>List enumeration\<close> |
9015 | 20 |
|
80786
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
21 |
nonterminal list_args |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
22 |
syntax |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
23 |
"" :: "'a \<Rightarrow> list_args" ("_") |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
24 |
"_list_args" :: "'a \<Rightarrow> list_args \<Rightarrow> list_args" ("_,/ _") |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
25 |
"_list" :: "list_args => 'a list" ("[(_)]") |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
26 |
syntax_consts |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
27 |
"_list_args" "_list" == Cons |
9015 | 28 |
translations |
80786
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
29 |
"[x, xs]" == "x#[xs]" |
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents:
69597
diff
changeset
|
30 |
"[x]" == "x#[]" |
9015 | 31 |
|
17311 | 32 |
typedecl person |
9015 | 33 |
|
51311 | 34 |
axiomatization |
35 |
append :: "['a list, 'a list, 'a list] => bool" and |
|
36 |
reverse :: "['a list, 'a list] => bool" and |
|
9015 | 37 |
|
51311 | 38 |
mappred :: "[('a => 'b => bool), 'a list, 'b list] => bool" and |
39 |
mapfun :: "[('a => 'b), 'a list, 'b list] => bool" and |
|
9015 | 40 |
|
51311 | 41 |
bob :: person and |
42 |
sue :: person and |
|
43 |
ned :: person and |
|
9015 | 44 |
|
51311 | 45 |
nat23 :: nat ("23") and |
46 |
nat24 :: nat ("24") and |
|
47 |
nat25 :: nat ("25") and |
|
9015 | 48 |
|
51311 | 49 |
age :: "[person, nat] => bool" and |
9015 | 50 |
|
51311 | 51 |
eq :: "['a, 'a] => bool" and |
9015 | 52 |
|
51311 | 53 |
empty :: "['b] => bool" and |
54 |
add :: "['a, 'b, 'b] => bool" and |
|
55 |
remove :: "['a, 'b, 'b] => bool" and |
|
17311 | 56 |
bag_appl:: "['a, 'a, 'a, 'a] => bool" |
51311 | 57 |
where |
58 |
append: "\<And>x xs ys zs. append [] xs xs .. |
|
59 |
append (x#xs) ys (x#zs) :- append xs ys zs" and |
|
67613 | 60 |
reverse: "\<And>L1 L2. reverse L1 L2 :- (\<forall>rev_aux. |
61 |
(\<forall>L. rev_aux [] L L ).. |
|
62 |
(\<forall>X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3)) |
|
51311 | 63 |
=> rev_aux L1 L2 [])" and |
9015 | 64 |
|
51311 | 65 |
mappred: "\<And>x xs y ys P. mappred P [] [] .. |
66 |
mappred P (x#xs) (y#ys) :- P x y & mappred P xs ys" and |
|
67 |
mapfun: "\<And>x xs ys f. mapfun f [] [] .. |
|
68 |
mapfun f (x#xs) (f x#ys) :- mapfun f xs ys" and |
|
9015 | 69 |
|
17311 | 70 |
age: "age bob 24 .. |
71 |
age sue 23 .. |
|
51311 | 72 |
age ned 23" and |
9015 | 73 |
|
51311 | 74 |
eq: "\<And>x. eq x x" and |
9015 | 75 |
|
76 |
(* actual definitions of empty and add is hidden -> yields abstract data type *) |
|
77 |
||
67613 | 78 |
bag_appl: "\<And>A B X Y. bag_appl A B X Y:- (\<exists>S1 S2 S3 S4 S5. |
17311 | 79 |
empty S1 & |
80 |
add A S1 S2 & |
|
81 |
add B S2 S3 & |
|
82 |
remove X S3 S4 & |
|
83 |
remove Y S4 S5 & |
|
84 |
empty S5)" |
|
85 |
||
21425 | 86 |
lemmas prog_Test = append reverse mappred mapfun age eq bag_appl |
87 |
||
61337 | 88 |
schematic_goal "append ?x ?y [a,b,c,d]" |
21425 | 89 |
apply (prolog prog_Test) |
90 |
back |
|
91 |
back |
|
92 |
back |
|
93 |
back |
|
94 |
done |
|
95 |
||
61337 | 96 |
schematic_goal "append [a,b] y ?L" |
21425 | 97 |
apply (prolog prog_Test) |
98 |
done |
|
99 |
||
67613 | 100 |
schematic_goal "\<forall>y. append [a,b] y (?L y)" |
21425 | 101 |
apply (prolog prog_Test) |
102 |
done |
|
103 |
||
61337 | 104 |
schematic_goal "reverse [] ?L" |
21425 | 105 |
apply (prolog prog_Test) |
106 |
done |
|
107 |
||
61337 | 108 |
schematic_goal "reverse [23] ?L" |
21425 | 109 |
apply (prolog prog_Test) |
110 |
done |
|
111 |
||
61337 | 112 |
schematic_goal "reverse [23,24,?x] ?L" |
21425 | 113 |
apply (prolog prog_Test) |
114 |
done |
|
115 |
||
61337 | 116 |
schematic_goal "reverse ?L [23,24,?x]" |
21425 | 117 |
apply (prolog prog_Test) |
118 |
done |
|
119 |
||
61337 | 120 |
schematic_goal "mappred age ?x [23,24]" |
21425 | 121 |
apply (prolog prog_Test) |
122 |
back |
|
123 |
done |
|
124 |
||
67613 | 125 |
schematic_goal "mappred (\<lambda>x y. \<exists>z. age z y) ?x [23,24]" |
21425 | 126 |
apply (prolog prog_Test) |
127 |
done |
|
128 |
||
61337 | 129 |
schematic_goal "mappred ?P [bob,sue] [24,23]" |
21425 | 130 |
apply (prolog prog_Test) |
131 |
done |
|
132 |
||
61337 | 133 |
schematic_goal "mapfun f [bob,bob,sue] [?x,?y,?z]" |
21425 | 134 |
apply (prolog prog_Test) |
135 |
done |
|
136 |
||
61337 | 137 |
schematic_goal "mapfun (%x. h x 25) [bob,sue] ?L" |
21425 | 138 |
apply (prolog prog_Test) |
139 |
done |
|
140 |
||
61337 | 141 |
schematic_goal "mapfun ?F [24,25] [h bob 24,h bob 25]" |
21425 | 142 |
apply (prolog prog_Test) |
143 |
done |
|
144 |
||
61337 | 145 |
schematic_goal "mapfun ?F [24] [h 24 24]" |
21425 | 146 |
apply (prolog prog_Test) |
147 |
back |
|
148 |
back |
|
149 |
back |
|
150 |
done |
|
151 |
||
152 |
lemma "True" |
|
153 |
apply (prolog prog_Test) |
|
154 |
done |
|
155 |
||
61337 | 156 |
schematic_goal "age ?x 24 & age ?y 23" |
21425 | 157 |
apply (prolog prog_Test) |
158 |
back |
|
159 |
done |
|
160 |
||
61337 | 161 |
schematic_goal "age ?x 24 | age ?x 23" |
21425 | 162 |
apply (prolog prog_Test) |
163 |
back |
|
164 |
back |
|
165 |
done |
|
166 |
||
67613 | 167 |
lemma "\<exists>x y. age x y" |
21425 | 168 |
apply (prolog prog_Test) |
169 |
done |
|
170 |
||
67613 | 171 |
lemma "\<forall>x y. append [] x x" |
21425 | 172 |
apply (prolog prog_Test) |
173 |
done |
|
174 |
||
61337 | 175 |
schematic_goal "age sue 24 .. age bob 23 => age ?x ?y" |
21425 | 176 |
apply (prolog prog_Test) |
177 |
back |
|
178 |
back |
|
179 |
back |
|
180 |
back |
|
181 |
done |
|
182 |
||
183 |
(*set trace_DEPTH_FIRST;*) |
|
184 |
lemma "age bob 25 :- age bob 24 => age bob 25" |
|
185 |
apply (prolog prog_Test) |
|
186 |
done |
|
187 |
(*reset trace_DEPTH_FIRST;*) |
|
188 |
||
67613 | 189 |
schematic_goal "(\<forall>x. age x 25 :- age x 23) => age ?x 25 & age ?y 25" |
21425 | 190 |
apply (prolog prog_Test) |
191 |
back |
|
192 |
back |
|
193 |
back |
|
194 |
done |
|
195 |
||
67613 | 196 |
lemma "\<forall>x. \<exists>y. eq x y" |
21425 | 197 |
apply (prolog prog_Test) |
198 |
done |
|
199 |
||
67613 | 200 |
schematic_goal "\<exists>P. P \<and> eq P ?x" |
21425 | 201 |
apply (prolog prog_Test) |
202 |
(* |
|
203 |
back |
|
204 |
back |
|
205 |
back |
|
206 |
back |
|
207 |
back |
|
208 |
back |
|
209 |
back |
|
210 |
back |
|
211 |
*) |
|
212 |
done |
|
213 |
||
67613 | 214 |
lemma "\<exists>P. eq P (True & True) \<and> P" |
21425 | 215 |
apply (prolog prog_Test) |
216 |
done |
|
217 |
||
67613 | 218 |
lemma "\<exists>P. eq P (\<or>) \<and> P k True" |
21425 | 219 |
apply (prolog prog_Test) |
220 |
done |
|
221 |
||
67613 | 222 |
lemma "\<exists>P. eq P (Q => True) \<and> P" |
21425 | 223 |
apply (prolog prog_Test) |
224 |
done |
|
225 |
||
226 |
(* P flexible: *) |
|
67613 | 227 |
lemma "(\<forall>P k l. P k l :- eq P Q) => Q a b" |
21425 | 228 |
apply (prolog prog_Test) |
229 |
done |
|
230 |
(* |
|
231 |
loops: |
|
232 |
lemma "(!P k l. P k l :- eq P (%x y. x | y)) => a | b" |
|
233 |
*) |
|
234 |
||
235 |
(* implication and disjunction in atom: *) |
|
67613 | 236 |
lemma "\<exists>Q. (\<forall>p q. R(q :- p) => R(Q p q)) \<and> Q (t | s) (s | t)" |
21425 | 237 |
by fast |
238 |
||
239 |
(* disjunction in atom: *) |
|
67613 | 240 |
lemma "(\<forall>P. g P :- (P => b \<or> a)) => g(a \<or> b)" |
69597 | 241 |
apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1") |
242 |
apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1") |
|
243 |
apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1") |
|
21425 | 244 |
prefer 2 |
245 |
apply fast |
|
246 |
apply fast |
|
247 |
done |
|
248 |
||
249 |
(* |
|
250 |
hangs: |
|
251 |
lemma "(!P. g P :- (P => b | a)) => g(a | b)" |
|
252 |
by fast |
|
253 |
*) |
|
254 |
||
67613 | 255 |
schematic_goal "\<forall>Emp Stk.( |
21425 | 256 |
empty (Emp::'b) .. |
67613 | 257 |
(\<forall>(X::nat) S. add X (S::'b) (Stk X S)) .. |
258 |
(\<forall>(X::nat) S. remove X ((Stk X S)::'b) S)) |
|
21425 | 259 |
=> bag_appl 23 24 ?X ?Y" |
260 |
oops |
|
261 |
||
67613 | 262 |
schematic_goal "\<forall>Qu. ( |
263 |
(\<forall>L. empty (Qu L L)) .. |
|
264 |
(\<forall>(X::nat) L K. add X (Qu L (X#K)) (Qu L K)) .. |
|
265 |
(\<forall>(X::nat) L K. remove X (Qu (X#L) K) (Qu L K))) |
|
21425 | 266 |
=> bag_appl 23 24 ?X ?Y" |
267 |
oops |
|
268 |
||
67613 | 269 |
lemma "D \<and> (\<forall>y. E) :- (\<forall>x. True \<and> True) :- True => D" |
21425 | 270 |
apply (prolog prog_Test) |
271 |
done |
|
272 |
||
61337 | 273 |
schematic_goal "P x .. P y => P ?X" |
21425 | 274 |
apply (prolog prog_Test) |
275 |
back |
|
276 |
done |
|
17311 | 277 |
|
9015 | 278 |
end |