| author | wenzelm | 
| Sun, 15 Apr 2007 14:31:44 +0200 | |
| changeset 22690 | 0b08f218f260 | 
| parent 16417 | 9bc16273c2d4 | 
| child 23733 | 3f8ad7418e55 | 
| permissions | -rw-r--r-- | 
| 10370 | 1 | (* ID: $Id$ *) | 
| 16417 | 2 | theory Advanced imports Even begin | 
| 10426 | 3 | |
| 10468 | 4 | |
| 5 | datatype 'f gterm = Apply 'f "'f gterm list" | |
| 6 | ||
| 7 | datatype integer_op = Number int | UnaryMinus | Plus; | |
| 8 | ||
| 9 | consts gterms :: "'f set \<Rightarrow> 'f gterm set" | |
| 10 | inductive "gterms F" | |
| 11 | intros | |
| 12 | step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> gterms F; f \<in> F\<rbrakk> | |
| 13 | \<Longrightarrow> (Apply f args) \<in> gterms F" | |
| 14 | ||
| 15 | lemma gterms_mono: "F\<subseteq>G \<Longrightarrow> gterms F \<subseteq> gterms G" | |
| 16 | apply clarify | |
| 17 | apply (erule gterms.induct) | |
| 10882 | 18 | txt{*
 | 
| 19 | @{subgoals[display,indent=0,margin=65]}
 | |
| 20 | *}; | |
| 10468 | 21 | apply blast | 
| 22 | done | |
| 23 | ||
| 10882 | 24 | |
| 10468 | 25 | text{*
 | 
| 26 | @{thm[display] even.cases[no_vars]}
 | |
| 27 | \rulename{even.cases}
 | |
| 28 | ||
| 29 | Just as a demo I include | |
| 30 | the two forms that Markus has made available. First the one for binding the | |
| 31 | result to a name | |
| 32 | ||
| 33 | *} | |
| 34 | ||
| 35 | inductive_cases Suc_Suc_cases [elim!]: | |
| 10426 | 36 | "Suc(Suc n) \<in> even" | 
| 37 | ||
| 10468 | 38 | thm Suc_Suc_cases; | 
| 39 | ||
| 40 | text{*
 | |
| 41 | @{thm[display] Suc_Suc_cases[no_vars]}
 | |
| 42 | \rulename{Suc_Suc_cases}
 | |
| 43 | ||
| 44 | and now the one for local usage: | |
| 45 | *} | |
| 46 | ||
| 47 | lemma "Suc(Suc n) \<in> even \<Longrightarrow> P n"; | |
| 48 | apply (ind_cases "Suc(Suc n) \<in> even"); | |
| 49 | oops | |
| 50 | ||
| 51 | inductive_cases gterm_Apply_elim [elim!]: "Apply f args \<in> gterms F" | |
| 52 | ||
| 53 | text{*this is what we get:
 | |
| 54 | ||
| 55 | @{thm[display] gterm_Apply_elim[no_vars]}
 | |
| 56 | \rulename{gterm_Apply_elim}
 | |
| 57 | *} | |
| 10426 | 58 | |
| 11173 | 59 | lemma gterms_IntI [rule_format, intro!]: | 
| 10468 | 60 | "t \<in> gterms F \<Longrightarrow> t \<in> gterms G \<longrightarrow> t \<in> gterms (F\<inter>G)" | 
| 61 | apply (erule gterms.induct) | |
| 10882 | 62 | txt{*
 | 
| 63 | @{subgoals[display,indent=0,margin=65]}
 | |
| 64 | *}; | |
| 10468 | 65 | apply blast | 
| 66 | done | |
| 67 | ||
| 68 | ||
| 69 | text{*
 | |
| 70 | @{thm[display] mono_Int[no_vars]}
 | |
| 71 | \rulename{mono_Int}
 | |
| 72 | *} | |
| 10426 | 73 | |
| 10468 | 74 | lemma gterms_Int_eq [simp]: | 
| 75 | "gterms (F\<inter>G) = gterms F \<inter> gterms G" | |
| 11173 | 76 | by (blast intro!: mono_Int monoI gterms_mono) | 
| 10468 | 77 | |
| 78 | ||
| 12156 
d2758965362e
new-style numerals without leading #, along with generic 0 and 1
 paulson parents: 
11711diff
changeset | 79 | text{*the following declaration isn't actually used*}
 | 
| 10468 | 80 | consts integer_arity :: "integer_op \<Rightarrow> nat" | 
| 81 | primrec | |
| 11711 | 82 | "integer_arity (Number n) = 0" | 
| 83 | "integer_arity UnaryMinus = 1" | |
| 84 | "integer_arity Plus = 2" | |
| 10468 | 85 | |
| 86 | consts well_formed_gterm :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set"
 | |
| 87 | inductive "well_formed_gterm arity" | |
| 88 | intros | |
| 89 | step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> well_formed_gterm arity; | |
| 90 | length args = arity f\<rbrakk> | |
| 91 | \<Longrightarrow> (Apply f args) \<in> well_formed_gterm arity" | |
| 92 | ||
| 10426 | 93 | |
| 10468 | 94 | consts well_formed_gterm' :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set"
 | 
| 95 | inductive "well_formed_gterm' arity" | |
| 96 | intros | |
| 97 | step[intro!]: "\<lbrakk>args \<in> lists (well_formed_gterm' arity); | |
| 98 | length args = arity f\<rbrakk> | |
| 99 | \<Longrightarrow> (Apply f args) \<in> well_formed_gterm' arity" | |
| 100 | monos lists_mono | |
| 101 | ||
| 102 | lemma "well_formed_gterm arity \<subseteq> well_formed_gterm' arity" | |
| 103 | apply clarify | |
| 10882 | 104 | txt{*
 | 
| 105 | The situation after clarify | |
| 106 | @{subgoals[display,indent=0,margin=65]}
 | |
| 107 | *}; | |
| 10468 | 108 | apply (erule well_formed_gterm.induct) | 
| 10882 | 109 | txt{*
 | 
| 110 | note the induction hypothesis! | |
| 111 | @{subgoals[display,indent=0,margin=65]}
 | |
| 112 | *}; | |
| 10468 | 113 | apply auto | 
| 114 | done | |
| 115 | ||
| 10426 | 116 | |
| 10370 | 117 | |
| 10468 | 118 | lemma "well_formed_gterm' arity \<subseteq> well_formed_gterm arity" | 
| 119 | apply clarify | |
| 10882 | 120 | txt{*
 | 
| 121 | The situation after clarify | |
| 122 | @{subgoals[display,indent=0,margin=65]}
 | |
| 123 | *}; | |
| 10468 | 124 | apply (erule well_formed_gterm'.induct) | 
| 10882 | 125 | txt{*
 | 
| 126 | note the induction hypothesis! | |
| 127 | @{subgoals[display,indent=0,margin=65]}
 | |
| 128 | *}; | |
| 10468 | 129 | apply auto | 
| 130 | done | |
| 10370 | 131 | |
| 132 | ||
| 10468 | 133 | text{*
 | 
| 134 | @{thm[display] lists_Int_eq[no_vars]}
 | |
| 135 | *} | |
| 136 | ||
| 137 | text{* the rest isn't used: too complicated.  OK for an exercise though.*}
 | |
| 138 | ||
| 139 | consts integer_signature :: "(integer_op * (unit list * unit)) set" | |
| 140 | inductive "integer_signature" | |
| 10370 | 141 | intros | 
| 10468 | 142 | Number: "(Number n, ([], ())) \<in> integer_signature" | 
| 143 | UnaryMinus: "(UnaryMinus, ([()], ())) \<in> integer_signature" | |
| 144 | Plus: "(Plus, ([(),()], ())) \<in> integer_signature" | |
| 145 | ||
| 10370 | 146 | |
| 10468 | 147 | consts well_typed_gterm :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set"
 | 
| 148 | inductive "well_typed_gterm sig" | |
| 10370 | 149 | intros | 
| 10468 | 150 | step[intro!]: | 
| 151 | "\<lbrakk>\<forall>pair \<in> set args. pair \<in> well_typed_gterm sig; | |
| 152 | sig f = (map snd args, rtype)\<rbrakk> | |
| 153 | \<Longrightarrow> (Apply f (map fst args), rtype) | |
| 154 | \<in> well_typed_gterm sig" | |
| 155 | ||
| 156 | consts well_typed_gterm' :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set"
 | |
| 157 | inductive "well_typed_gterm' sig" | |
| 158 | intros | |
| 159 | step[intro!]: | |
| 160 | "\<lbrakk>args \<in> lists(well_typed_gterm' sig); | |
| 161 | sig f = (map snd args, rtype)\<rbrakk> | |
| 162 | \<Longrightarrow> (Apply f (map fst args), rtype) | |
| 163 | \<in> well_typed_gterm' sig" | |
| 10370 | 164 | monos lists_mono | 
| 165 | ||
| 166 | ||
| 10468 | 167 | lemma "well_typed_gterm sig \<subseteq> well_typed_gterm' sig" | 
| 10370 | 168 | apply clarify | 
| 10468 | 169 | apply (erule well_typed_gterm.induct) | 
| 10370 | 170 | apply auto | 
| 171 | done | |
| 172 | ||
| 10468 | 173 | lemma "well_typed_gterm' sig \<subseteq> well_typed_gterm sig" | 
| 10370 | 174 | apply clarify | 
| 10468 | 175 | apply (erule well_typed_gterm'.induct) | 
| 10370 | 176 | apply auto | 
| 177 | done | |
| 178 | ||
| 10468 | 179 | |
| 10370 | 180 | end | 
| 181 |