| author | blanchet | 
| Fri, 30 May 2014 12:27:51 +0200 | |
| changeset 57124 | e4c2c792226f | 
| parent 54555 | e8c5e95d338b | 
| child 58814 | 4c0ad4162cb7 | 
| permissions | -rw-r--r-- | 
| 54540 | 1  | 
(* Title: HOL/Coinduction.thy  | 
| 54026 | 2  | 
Author: Johannes Hölzl, TU Muenchen  | 
3  | 
Author: Dmitriy Traytel, TU Muenchen  | 
|
4  | 
Copyright 2013  | 
|
5  | 
||
6  | 
Coinduction method that avoids some boilerplate compared to coinduct.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
header {* Coinduction Method *}
 | 
|
10  | 
||
11  | 
theory Coinduction  | 
|
| 54555 | 12  | 
imports Ctr_Sugar  | 
| 54026 | 13  | 
begin  | 
14  | 
||
15  | 
ML_file "Tools/coinduction.ML"  | 
|
16  | 
||
17  | 
setup Coinduction.setup  | 
|
18  | 
||
19  | 
end  |