author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 0 | 7949f97df77a |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: Substitutions/alist.ML |
2 |
Author: Martin Coen, Cambridge University Computer Laboratory |
|
3 |
Copyright 1993 University of Cambridge |
|
4 |
||
5 |
For alist.thy. |
|
6 |
*) |
|
7 |
||
8 |
open AList; |
|
9 |
||
10 |
(*********) |
|
11 |
||
12 |
val al_defs = [alist_rec_def,assoc_def]; |
|
13 |
||
14 |
val alist_ss = pair_ss addsimps [Nil_not_Cons,Cons_not_Nil,Cons_Cons_eq, |
|
15 |
list_rec_Nil,list_rec_Cons]; |
|
16 |
||
17 |
val al_rews = |
|
18 |
let fun mk_thm s = prove_goalw AList.thy al_defs s |
|
19 |
(fn _ => [simp_tac alist_ss 1]) |
|
20 |
in map mk_thm |
|
21 |
["alist_rec(Nil,c,d) = c", |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
22 |
"alist_rec(<a,b>#al,c,d) = d(a,b,al,alist_rec(al,c,d))", |
0 | 23 |
"assoc(v,d,Nil) = d", |
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
24 |
"assoc(v,d,<a,b>#al) = if(v=a,b,assoc(v,d,al))"] end; |
0 | 25 |
|
26 |
(*********) |
|
27 |
||
28 |
val prems = goal AList.thy |
|
29 |
"[| P(Nil); \ |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
30 |
\ !!x y xs. P(xs) ==> P(<x,y>#xs) |] ==> P(l)"; |
0 | 31 |
by (list_ind_tac "l" 1); |
32 |
by (resolve_tac prems 1); |
|
33 |
by (rtac PairE 1); |
|
34 |
by (etac ssubst 1); |
|
35 |
by (resolve_tac prems 1); |
|
36 |
by (assume_tac 1); |
|
37 |
val alist_induct = result(); |
|
38 |
||
39 |
(*Perform induction on xs. *) |
|
40 |
fun alist_ind_tac a M = |
|
41 |
EVERY [res_inst_tac [("l",a)] alist_induct M, |
|
42 |
rename_last_tac a ["1"] (M+1)]; |
|
43 |
||
44 |
(* |
|
45 |
val prems = goal AList.thy |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
0
diff
changeset
|
46 |
"[| P(Nil); !! x y xs. P(xs) --> P(<x,y>#xs) |] ==> P(a)"; |
0 | 47 |
by (alist_ind_tac "a" 1); |
48 |
by (ALLGOALS (cut_facts_tac prems THEN' fast_tac HOL_cs)); |
|
49 |
val alist_induct2 = result(); |
|
50 |
||
51 |
add_inds alist_induct2; |
|
52 |
*) |