blanchet@54397: (* Title: HOL/Ctr_Sugar.thy blanchet@49486: Author: Jasmin Blanchette, TU Muenchen blanchet@55160: Author: Dmitriy Traytel, TU Muenchen blanchet@54397: Copyright 2012, 2013 blanchet@49075: blanchet@51797: Wrapping existing freely generated type's constructors. blanchet@49075: *) blanchet@49075: wenzelm@58889: section {* Wrapping Existing Freely Generated Type's Constructors *} blanchet@49075: blanchet@54006: theory Ctr_Sugar blanchet@54398: imports HOL blanchet@49075: keywords blanchet@54398: "print_case_translations" :: diag and blanchet@55468: "free_constructors" :: thy_goal blanchet@49075: begin blanchet@49075: blanchet@54398: consts blanchet@54398: case_guard :: "bool \ 'a \ ('a \ 'b) \ 'b" blanchet@54398: case_nil :: "'a \ 'b" blanchet@54398: case_cons :: "('a \ 'b) \ ('a \ 'b) \ 'a \ 'b" blanchet@54398: case_elem :: "'a \ 'b \ 'a \ 'b" blanchet@54398: case_abs :: "('c \ 'b) \ 'b" blanchet@57631: blanchet@54398: declare [[coercion_args case_guard - + -]] blanchet@54398: declare [[coercion_args case_cons - -]] blanchet@54398: declare [[coercion_args case_abs -]] blanchet@54398: declare [[coercion_args case_elem - +]] blanchet@54398: blanchet@54701: ML_file "Tools/Ctr_Sugar/case_translation.ML" blanchet@54398: blanchet@49312: lemma iffI_np: "\x \ \ y; \ x \ y\ \ \ x \ y" blanchet@58188: by (erule iffI) (erule contrapos_pn) blanchet@49312: blanchet@49486: lemma iff_contradict: blanchet@58188: "\ P \ P \ Q \ Q \ R" blanchet@58188: "\ Q \ P \ Q \ P \ R" blanchet@58188: by blast+ blanchet@49486: blanchet@54701: ML_file "Tools/Ctr_Sugar/ctr_sugar_util.ML" blanchet@54701: ML_file "Tools/Ctr_Sugar/ctr_sugar_tactics.ML" blanchet@54701: ML_file "Tools/Ctr_Sugar/ctr_sugar_code.ML" blanchet@54701: ML_file "Tools/Ctr_Sugar/ctr_sugar.ML" blanchet@49309: blanchet@49075: end