| author | hoelzl | 
| Thu, 13 Nov 2014 17:19:52 +0100 | |
| changeset 59000 | 6eb0725503fc | 
| parent 58372 | bfd497f2f4c2 | 
| child 61943 | 7fba644ed827 | 
| permissions | -rw-r--r-- | 
| 8840 | 1  | 
(* Title: HOL/Induct/Sexp.thy  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 1992 University of Cambridge  | 
|
4  | 
||
5  | 
S-expressions, general binary trees for defining recursive data  | 
|
6  | 
structures by hand.  | 
|
7  | 
*)  | 
|
8  | 
||
| 55017 | 9  | 
theory Sexp  | 
| 
58372
 
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
 
blanchet 
parents: 
58112 
diff
changeset
 | 
10  | 
imports "~~/src/HOL/Library/Old_Datatype"  | 
| 55017 | 11  | 
begin  | 
| 20801 | 12  | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55017 
diff
changeset
 | 
13  | 
type_synonym 'a item = "'a Old_Datatype.item"  | 
| 
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55017 
diff
changeset
 | 
14  | 
abbreviation "Leaf == Old_Datatype.Leaf"  | 
| 
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
55017 
diff
changeset
 | 
15  | 
abbreviation "Numb == Old_Datatype.Numb"  | 
| 20801 | 16  | 
|
| 23746 | 17  | 
inductive_set  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
18  | 
sexp :: "'a item set"  | 
| 23746 | 19  | 
where  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
20  | 
LeafI: "Leaf(a) \<in> sexp"  | 
| 23746 | 21  | 
| NumbI: "Numb(i) \<in> sexp"  | 
22  | 
| SconsI: "[| M \<in> sexp; N \<in> sexp |] ==> Scons M N \<in> sexp"  | 
|
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
23  | 
|
| 19736 | 24  | 
definition  | 
| 8840 | 25  | 
sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b,  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20820 
diff
changeset
 | 
26  | 
'a item] => 'b" where  | 
| 20801 | 27  | 
"sexp_case c d e M = (THE z. (EX x. M=Leaf(x) & z=c(x))  | 
28  | 
| (EX k. M=Numb(k) & z=d(k))  | 
|
29  | 
| (EX N1 N2. M = Scons N1 N2 & z=e N1 N2))"  | 
|
| 8840 | 30  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20820 
diff
changeset
 | 
31  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20820 
diff
changeset
 | 
32  | 
  pred_sexp :: "('a item * 'a item)set" where
 | 
| 19736 | 33  | 
     "pred_sexp = (\<Union>M \<in> sexp. \<Union>N \<in> sexp. {(M, Scons M N), (N, Scons M N)})"
 | 
| 8840 | 34  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20820 
diff
changeset
 | 
35  | 
definition  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
36  | 
sexp_rec :: "['a item, 'a=>'b, nat=>'b,  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20820 
diff
changeset
 | 
37  | 
['a item, 'a item, 'b, 'b]=>'b] => 'b" where  | 
| 19736 | 38  | 
"sexp_rec M c d e = wfrec pred_sexp  | 
| 8840 | 39  | 
(%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2))) M"  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
40  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
41  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
42  | 
(** sexp_case **)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
43  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
44  | 
lemma sexp_case_Leaf [simp]: "sexp_case c d e (Leaf a) = c(a)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
45  | 
by (simp add: sexp_case_def, blast)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
46  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
47  | 
lemma sexp_case_Numb [simp]: "sexp_case c d e (Numb k) = d(k)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
48  | 
by (simp add: sexp_case_def, blast)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
49  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
50  | 
lemma sexp_case_Scons [simp]: "sexp_case c d e (Scons M N) = e M N"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
51  | 
by (simp add: sexp_case_def)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
52  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
53  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
54  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
55  | 
(** Introduction rules for sexp constructors **)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
56  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
57  | 
lemma sexp_In0I: "M \<in> sexp ==> In0(M) \<in> sexp"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
58  | 
apply (simp add: In0_def)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
59  | 
apply (erule sexp.NumbI [THEN sexp.SconsI])  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
60  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
61  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
62  | 
lemma sexp_In1I: "M \<in> sexp ==> In1(M) \<in> sexp"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
63  | 
apply (simp add: In1_def)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
64  | 
apply (erule sexp.NumbI [THEN sexp.SconsI])  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
65  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
66  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
67  | 
declare sexp.intros [intro,simp]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
68  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
69  | 
lemma range_Leaf_subset_sexp: "range(Leaf) <= sexp"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
70  | 
by blast  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
71  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
72  | 
lemma Scons_D: "Scons M N \<in> sexp ==> M \<in> sexp & N \<in> sexp"  | 
| 18413 | 73  | 
by (induct S == "Scons M N" set: sexp) auto  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
74  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
75  | 
(** Introduction rules for 'pred_sexp' **)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
76  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
77  | 
lemma pred_sexp_subset_Sigma: "pred_sexp <= sexp <*> sexp"  | 
| 44918 | 78  | 
by (simp add: pred_sexp_def) blast  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
79  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
80  | 
(* (a,b) \<in> pred_sexp^+ ==> a \<in> sexp *)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
81  | 
lemmas trancl_pred_sexpD1 =  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
82  | 
pred_sexp_subset_Sigma  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
23746 
diff
changeset
 | 
83  | 
[THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD1]  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
84  | 
and trancl_pred_sexpD2 =  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
85  | 
pred_sexp_subset_Sigma  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
23746 
diff
changeset
 | 
86  | 
[THEN trancl_subset_Sigma, THEN subsetD, THEN SigmaD2]  | 
| 
13079
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
87  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
88  | 
lemma pred_sexpI1:  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
89  | 
"[| M \<in> sexp; N \<in> sexp |] ==> (M, Scons M N) \<in> pred_sexp"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
90  | 
by (simp add: pred_sexp_def, blast)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
91  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
92  | 
lemma pred_sexpI2:  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
93  | 
"[| M \<in> sexp; N \<in> sexp |] ==> (N, Scons M N) \<in> pred_sexp"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
94  | 
by (simp add: pred_sexp_def, blast)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
95  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
96  | 
(*Combinations involving transitivity and the rules above*)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
97  | 
lemmas pred_sexp_t1 [simp] = pred_sexpI1 [THEN r_into_trancl]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
98  | 
and pred_sexp_t2 [simp] = pred_sexpI2 [THEN r_into_trancl]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
99  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
100  | 
lemmas pred_sexp_trans1 [simp] = trans_trancl [THEN transD, OF _ pred_sexp_t1]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
101  | 
and pred_sexp_trans2 [simp] = trans_trancl [THEN transD, OF _ pred_sexp_t2]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
102  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
103  | 
(*Proves goals of the form (M,N):pred_sexp^+ provided M,N:sexp*)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
104  | 
declare cut_apply [simp]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
105  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
106  | 
lemma pred_sexpE:  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
107  | 
"[| p \<in> pred_sexp;  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
108  | 
!!M N. [| p = (M, Scons M N); M \<in> sexp; N \<in> sexp |] ==> R;  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
109  | 
!!M N. [| p = (N, Scons M N); M \<in> sexp; N \<in> sexp |] ==> R  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
110  | 
|] ==> R"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
111  | 
by (simp add: pred_sexp_def, blast)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
112  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
113  | 
lemma wf_pred_sexp: "wf(pred_sexp)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
114  | 
apply (rule pred_sexp_subset_Sigma [THEN wfI])  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
115  | 
apply (erule sexp.induct)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
116  | 
apply (blast elim!: pred_sexpE)+  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
117  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
118  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
119  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
120  | 
(*** sexp_rec -- by wf recursion on pred_sexp ***)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
121  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
122  | 
lemma sexp_rec_unfold_lemma:  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
123  | 
"(%M. sexp_rec M c d e) ==  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
124  | 
wfrec pred_sexp (%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2)))"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
125  | 
by (simp add: sexp_rec_def)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
126  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
127  | 
lemmas sexp_rec_unfold = def_wfrec [OF sexp_rec_unfold_lemma wf_pred_sexp]  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
128  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
129  | 
(* sexp_rec a c d e =  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
130  | 
sexp_case c d  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
131  | 
(%N1 N2.  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
132  | 
e N1 N2 (cut (%M. sexp_rec M c d e) pred_sexp a N1)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
133  | 
(cut (%M. sexp_rec M c d e) pred_sexp a N2)) a  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
134  | 
*)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
135  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
136  | 
(** conversion rules **)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
137  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
138  | 
lemma sexp_rec_Leaf: "sexp_rec (Leaf a) c d h = c(a)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
139  | 
apply (subst sexp_rec_unfold)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
140  | 
apply (rule sexp_case_Leaf)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
141  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
142  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
143  | 
lemma sexp_rec_Numb: "sexp_rec (Numb k) c d h = d(k)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
144  | 
apply (subst sexp_rec_unfold)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
145  | 
apply (rule sexp_case_Numb)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
146  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
147  | 
|
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
148  | 
lemma sexp_rec_Scons: "[| M \<in> sexp; N \<in> sexp |] ==>  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
149  | 
sexp_rec (Scons M N) c d h = h M N (sexp_rec M c d h) (sexp_rec N c d h)"  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
150  | 
apply (rule sexp_rec_unfold [THEN trans])  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
151  | 
apply (simp add: pred_sexpI1 pred_sexpI2)  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
152  | 
done  | 
| 
 
e7738aa7267f
conversion of Induct/{Slist,Sexp} to Isar scripts
 
paulson 
parents: 
11481 
diff
changeset
 | 
153  | 
|
| 8840 | 154  | 
end  |