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