| author | paulson | 
| Wed, 09 Feb 2000 11:45:10 +0100 | |
| changeset 8216 | e4b3192dfefa | 
| parent 5440 | f099dffd0f18 | 
| child 8703 | 816d8f6513be | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/Sexp | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 923 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 6 | S-expressions, general binary trees for defining recursive data structures | |
| 7 | *) | |
| 8 | ||
| 9 | open Sexp; | |
| 10 | ||
| 11 | (** sexp_case **) | |
| 12 | ||
| 5069 | 13 | Goalw [sexp_case_def] "sexp_case c d e (Leaf a) = c(a)"; | 
| 4535 | 14 | by (Blast_tac 1); | 
| 923 | 15 | qed "sexp_case_Leaf"; | 
| 16 | ||
| 5069 | 17 | Goalw [sexp_case_def] "sexp_case c d e (Numb k) = d(k)"; | 
| 4535 | 18 | by (Blast_tac 1); | 
| 923 | 19 | qed "sexp_case_Numb"; | 
| 20 | ||
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 21 | Goalw [sexp_case_def] "sexp_case c d e (Scons M N) = e M N"; | 
| 4535 | 22 | by (Blast_tac 1); | 
| 923 | 23 | qed "sexp_case_Scons"; | 
| 24 | ||
| 4521 | 25 | Addsimps [sexp_case_Leaf, sexp_case_Numb, sexp_case_Scons]; | 
| 26 | ||
| 923 | 27 | |
| 28 | (** Introduction rules for sexp constructors **) | |
| 29 | ||
| 3029 | 30 | val [prem] = goalw Sexp.thy [In0_def] "M: sexp ==> In0(M) : sexp"; | 
| 923 | 31 | by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1); | 
| 32 | qed "sexp_In0I"; | |
| 33 | ||
| 3029 | 34 | val [prem] = goalw Sexp.thy [In1_def] "M: sexp ==> In1(M) : sexp"; | 
| 923 | 35 | by (rtac (prem RS (sexp.NumbI RS sexp.SconsI)) 1); | 
| 36 | qed "sexp_In1I"; | |
| 37 | ||
| 5440 | 38 | AddIs sexp.intrs; | 
| 1760 
6f41a494f3b1
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1642diff
changeset | 39 | |
| 5069 | 40 | Goal "range(Leaf) <= sexp"; | 
| 2892 | 41 | by (Blast_tac 1); | 
| 923 | 42 | qed "range_Leaf_subset_sexp"; | 
| 43 | ||
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 44 | val [major] = goal Sexp.thy "Scons M N : sexp ==> M: sexp & N: sexp"; | 
| 923 | 45 | by (rtac (major RS setup_induction) 1); | 
| 46 | by (etac sexp.induct 1); | |
| 2892 | 47 | by (ALLGOALS Blast_tac); | 
| 923 | 48 | qed "Scons_D"; | 
| 49 | ||
| 50 | (** Introduction rules for 'pred_sexp' **) | |
| 51 | ||
| 5069 | 52 | Goalw [pred_sexp_def] "pred_sexp <= sexp Times sexp"; | 
| 2892 | 53 | by (Blast_tac 1); | 
| 923 | 54 | qed "pred_sexp_subset_Sigma"; | 
| 55 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 56 | (* (a,b) : pred_sexp^+ ==> a : sexp *) | 
| 923 | 57 | val trancl_pred_sexpD1 = | 
| 58 | pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD1 | |
| 59 | and trancl_pred_sexpD2 = | |
| 60 | pred_sexp_subset_Sigma RS trancl_subset_Sigma RS subsetD RS SigmaD2; | |
| 61 | ||
| 5069 | 62 | Goalw [pred_sexp_def] | 
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 63 | "!!M. [| M: sexp; N: sexp |] ==> (M, Scons M N) : pred_sexp"; | 
| 2892 | 64 | by (Blast_tac 1); | 
| 923 | 65 | qed "pred_sexpI1"; | 
| 66 | ||
| 5069 | 67 | Goalw [pred_sexp_def] | 
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 68 | "!!M. [| M: sexp; N: sexp |] ==> (N, Scons M N) : pred_sexp"; | 
| 2892 | 69 | by (Blast_tac 1); | 
| 923 | 70 | qed "pred_sexpI2"; | 
| 71 | ||
| 72 | (*Combinations involving transitivity and the rules above*) | |
| 73 | val pred_sexp_t1 = pred_sexpI1 RS r_into_trancl | |
| 74 | and pred_sexp_t2 = pred_sexpI2 RS r_into_trancl; | |
| 75 | ||
| 76 | val pred_sexp_trans1 = pred_sexp_t1 RSN (2, trans_trancl RS transD) | |
| 77 | and pred_sexp_trans2 = pred_sexp_t2 RSN (2, trans_trancl RS transD); | |
| 78 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
923diff
changeset | 79 | (*Proves goals of the form (M,N):pred_sexp^+ provided M,N:sexp*) | 
| 1264 | 80 | Addsimps (sexp.intrs @ [pred_sexp_t1, pred_sexp_t2, | 
| 1465 | 81 | pred_sexp_trans1, pred_sexp_trans2, cut_apply]); | 
| 923 | 82 | |
| 83 | val major::prems = goalw Sexp.thy [pred_sexp_def] | |
| 3029 | 84 | "[| p : pred_sexp; \ | 
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 85 | \ !!M N. [| p = (M, Scons M N); M: sexp; N: sexp |] ==> R; \ | 
| 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 86 | \ !!M N. [| p = (N, Scons M N); M: sexp; N: sexp |] ==> R \ | 
| 923 | 87 | \ |] ==> R"; | 
| 88 | by (cut_facts_tac [major] 1); | |
| 89 | by (REPEAT (eresolve_tac ([asm_rl,emptyE,insertE,UN_E]@prems) 1)); | |
| 90 | qed "pred_sexpE"; | |
| 91 | ||
| 5069 | 92 | Goal "wf(pred_sexp)"; | 
| 923 | 93 | by (rtac (pred_sexp_subset_Sigma RS wfI) 1); | 
| 94 | by (etac sexp.induct 1); | |
| 4089 | 95 | by (ALLGOALS (blast_tac (claset() addSEs [allE, pred_sexpE]))); | 
| 923 | 96 | qed "wf_pred_sexp"; | 
| 97 | ||
| 3029 | 98 | |
| 923 | 99 | (*** sexp_rec -- by wf recursion on pred_sexp ***) | 
| 100 | ||
| 5278 | 101 | Goal "(%M. sexp_rec M c d e) = wfrec pred_sexp \ | 
| 1475 | 102 | \ (%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2)))"; | 
| 103 | by (simp_tac (HOL_ss addsimps [sexp_rec_def]) 1); | |
| 4521 | 104 | |
| 105 | (* sexp_rec a c d e = | |
| 106 | sexp_case c d | |
| 107 | (%N1 N2. | |
| 108 | e N1 N2 (cut (%M. sexp_rec M c d e) pred_sexp a N1) | |
| 109 | (cut (%M. sexp_rec M c d e) pred_sexp a N2)) a | |
| 110 | *) | |
| 3029 | 111 | bind_thm("sexp_rec_unfold", 
 | 
| 112 | [result() RS eq_reflection, wf_pred_sexp] MRS def_wfrec); | |
| 923 | 113 | |
| 3029 | 114 | (** conversion rules **) | 
| 923 | 115 | |
| 5069 | 116 | Goal "sexp_rec (Leaf a) c d h = c(a)"; | 
| 923 | 117 | by (stac sexp_rec_unfold 1); | 
| 118 | by (rtac sexp_case_Leaf 1); | |
| 119 | qed "sexp_rec_Leaf"; | |
| 120 | ||
| 5069 | 121 | Goal "sexp_rec (Numb k) c d h = d(k)"; | 
| 923 | 122 | by (stac sexp_rec_unfold 1); | 
| 123 | by (rtac sexp_case_Numb 1); | |
| 124 | qed "sexp_rec_Numb"; | |
| 125 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 126 | Goal "[| M: sexp; N: sexp |] ==> \ | 
| 5191 
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
 berghofe parents: 
5143diff
changeset | 127 | \ sexp_rec (Scons M N) c d h = h M N (sexp_rec M c d h) (sexp_rec N c d h)"; | 
| 923 | 128 | by (rtac (sexp_rec_unfold RS trans) 1); | 
| 4521 | 129 | by (asm_simp_tac (simpset() addsimps [pred_sexpI1, pred_sexpI2]) 1); | 
| 923 | 130 | qed "sexp_rec_Scons"; |