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