src/HOL/Isar_examples/ExprCompiler.thy
changeset 6480 c58bc3d2ba0f
parent 6444 2ebe9e630cab
child 6503 914729515c26
equal deleted inserted replaced
6479:b0448cab1b1e 6480:c58bc3d2ba0f
   103 
   103 
   104 ML "reset HOL_quantifiers";
   104 ML "reset HOL_quantifiers";
   105 
   105 
   106 lemma exec_append:
   106 lemma exec_append:
   107   "ALL stack. exec (xs @ ys) stack env = exec ys (exec xs stack env) env" (is "??P xs");
   107   "ALL stack. exec (xs @ ys) stack env = exec ys (exec xs stack env) env" (is "??P xs");
   108 proof ("induct" ??P xs type: list);
   108 proof (induct ??P xs type: list);
   109   show "??P []"; by simp;
   109   show "??P []"; by simp;
   110 
   110 
   111   fix x xs;
   111   fix x xs;
   112   assume "??P xs";
   112   assume "??P xs";
   113   show "??P (x # xs)" (is "??Q x");
   113   show "??P (x # xs)" (is "??Q x");
   114   proof ("induct" ??Q x type: instr);
   114   proof (induct ??Q x type: instr);
   115     fix val; show "??Q (Const val)"; by brute_force;
   115     fix val; show "??Q (Const val)"; by brute_force;
   116     next;
   116     next;
   117     fix adr; show "??Q (Load adr)"; by brute_force;
   117     fix adr; show "??Q (Load adr)"; by brute_force;
   118     next;
   118     next;
   119     fix fun; show "??Q (Apply fun)"; by brute_force;
   119     fix fun; show "??Q (Apply fun)"; by brute_force;
   121 qed;
   121 qed;
   122 
   122 
   123 
   123 
   124 lemma exec_comp:
   124 lemma exec_comp:
   125   "ALL stack. exec (comp e) stack env = eval e env # stack" (is "??P e");
   125   "ALL stack. exec (comp e) stack env = eval e env # stack" (is "??P e");
   126 proof ("induct" ??P e type: expr);
   126 proof (induct ??P e type: expr);
   127 
   127 
   128   fix adr; show "??P (Variable adr)"; by brute_force;
   128   fix adr; show "??P (Variable adr)"; by brute_force;
   129   next;
   129   next;
   130   fix val; show "??P (Constant val)"; by brute_force;
   130   fix val; show "??P (Constant val)"; by brute_force;
   131   next;
   131   next;