| 10370 |      1 | (* ID:         $Id$ *)
 | 
| 10426 |      2 | theory Advanced = Even:
 | 
|  |      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 | 
 | 
|  |     79 | consts integer_arity :: "integer_op \<Rightarrow> nat"
 | 
|  |     80 | primrec
 | 
|  |     81 | "integer_arity (Number n)        = #0"
 | 
|  |     82 | "integer_arity UnaryMinus        = #1"
 | 
|  |     83 | "integer_arity Plus              = #2"
 | 
|  |     84 | 
 | 
|  |     85 | consts well_formed_gterm :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set"
 | 
|  |     86 | inductive "well_formed_gterm arity"
 | 
|  |     87 | intros
 | 
|  |     88 | step[intro!]: "\<lbrakk>\<forall>t \<in> set args. t \<in> well_formed_gterm arity;  
 | 
|  |     89 |                 length args = arity f\<rbrakk>
 | 
|  |     90 |                \<Longrightarrow> (Apply f args) \<in> well_formed_gterm arity"
 | 
|  |     91 | 
 | 
| 10426 |     92 | 
 | 
| 10468 |     93 | consts well_formed_gterm' :: "('f \<Rightarrow> nat) \<Rightarrow> 'f gterm set"
 | 
|  |     94 | inductive "well_formed_gterm' arity"
 | 
|  |     95 | intros
 | 
|  |     96 | step[intro!]: "\<lbrakk>args \<in> lists (well_formed_gterm' arity);  
 | 
|  |     97 |                 length args = arity f\<rbrakk>
 | 
|  |     98 |                \<Longrightarrow> (Apply f args) \<in> well_formed_gterm' arity"
 | 
|  |     99 | monos lists_mono
 | 
|  |    100 | 
 | 
|  |    101 | lemma "well_formed_gterm arity \<subseteq> well_formed_gterm' arity"
 | 
|  |    102 | apply clarify
 | 
| 10882 |    103 | txt{*
 | 
|  |    104 | The situation after clarify
 | 
|  |    105 | @{subgoals[display,indent=0,margin=65]}
 | 
|  |    106 | *};
 | 
| 10468 |    107 | apply (erule well_formed_gterm.induct)
 | 
| 10882 |    108 | txt{*
 | 
|  |    109 | note the induction hypothesis!
 | 
|  |    110 | @{subgoals[display,indent=0,margin=65]}
 | 
|  |    111 | *};
 | 
| 10468 |    112 | apply auto
 | 
|  |    113 | done
 | 
|  |    114 | 
 | 
| 10426 |    115 | 
 | 
| 10370 |    116 | 
 | 
| 10468 |    117 | lemma "well_formed_gterm' arity \<subseteq> well_formed_gterm arity"
 | 
|  |    118 | apply clarify
 | 
| 10882 |    119 | txt{*
 | 
|  |    120 | The situation after clarify
 | 
|  |    121 | @{subgoals[display,indent=0,margin=65]}
 | 
|  |    122 | *};
 | 
| 10468 |    123 | apply (erule well_formed_gterm'.induct)
 | 
| 10882 |    124 | txt{*
 | 
|  |    125 | note the induction hypothesis!
 | 
|  |    126 | @{subgoals[display,indent=0,margin=65]}
 | 
|  |    127 | *};
 | 
| 10468 |    128 | apply auto
 | 
|  |    129 | done
 | 
| 10370 |    130 | 
 | 
|  |    131 | 
 | 
| 10468 |    132 | text{*
 | 
|  |    133 | @{thm[display] lists_Int_eq[no_vars]}
 | 
|  |    134 | *}
 | 
|  |    135 | 
 | 
|  |    136 | text{* the rest isn't used: too complicated.  OK for an exercise though.*}
 | 
|  |    137 | 
 | 
|  |    138 | consts integer_signature :: "(integer_op * (unit list * unit)) set"
 | 
|  |    139 | inductive "integer_signature"
 | 
| 10370 |    140 | intros
 | 
| 10468 |    141 | Number:     "(Number n,   ([], ())) \<in> integer_signature"
 | 
|  |    142 | UnaryMinus: "(UnaryMinus, ([()], ())) \<in> integer_signature"
 | 
|  |    143 | Plus:       "(Plus,       ([(),()], ())) \<in> integer_signature"
 | 
|  |    144 | 
 | 
| 10370 |    145 | 
 | 
| 10468 |    146 | consts well_typed_gterm :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set"
 | 
|  |    147 | inductive "well_typed_gterm sig"
 | 
| 10370 |    148 | intros
 | 
| 10468 |    149 | step[intro!]: 
 | 
|  |    150 |     "\<lbrakk>\<forall>pair \<in> set args. pair \<in> well_typed_gterm sig; 
 | 
|  |    151 |       sig f = (map snd args, rtype)\<rbrakk>
 | 
|  |    152 |      \<Longrightarrow> (Apply f (map fst args), rtype) 
 | 
|  |    153 |          \<in> well_typed_gterm sig"
 | 
|  |    154 | 
 | 
|  |    155 | consts well_typed_gterm' :: "('f \<Rightarrow> 't list * 't) \<Rightarrow> ('f gterm * 't)set"
 | 
|  |    156 | inductive "well_typed_gterm' sig"
 | 
|  |    157 | intros
 | 
|  |    158 | step[intro!]: 
 | 
|  |    159 |     "\<lbrakk>args \<in> lists(well_typed_gterm' sig); 
 | 
|  |    160 |       sig f = (map snd args, rtype)\<rbrakk>
 | 
|  |    161 |      \<Longrightarrow> (Apply f (map fst args), rtype) 
 | 
|  |    162 |          \<in> well_typed_gterm' sig"
 | 
| 10370 |    163 | monos lists_mono
 | 
|  |    164 | 
 | 
|  |    165 | 
 | 
| 10468 |    166 | lemma "well_typed_gterm sig \<subseteq> well_typed_gterm' sig"
 | 
| 10370 |    167 | apply clarify
 | 
| 10468 |    168 | apply (erule well_typed_gterm.induct)
 | 
| 10370 |    169 | apply auto
 | 
|  |    170 | done
 | 
|  |    171 | 
 | 
| 10468 |    172 | lemma "well_typed_gterm' sig \<subseteq> well_typed_gterm sig"
 | 
| 10370 |    173 | apply clarify
 | 
| 10468 |    174 | apply (erule well_typed_gterm'.induct)
 | 
| 10370 |    175 | apply auto
 | 
|  |    176 | done
 | 
|  |    177 | 
 | 
| 10468 |    178 | 
 | 
| 10370 |    179 | end
 | 
|  |    180 | 
 |