src/HOL/Ctr_Sugar.thy
author nipkow
Tue, 05 Nov 2019 21:07:03 +0100
changeset 71044 cb504351d058
parent 69605 a96320074298
permissions -rw-r--r--
tuned

(*  Title:      HOL/Ctr_Sugar.thy
    Author:     Jasmin Blanchette, TU Muenchen
    Author:     Dmitriy Traytel, TU Muenchen
    Copyright   2012, 2013

Wrapping existing freely generated type's constructors.
*)

section \<open>Wrapping Existing Freely Generated Type's Constructors\<close>

theory Ctr_Sugar
imports HOL
keywords
  "print_case_translations" :: diag and
  "free_constructors" :: thy_goal
begin

consts
  case_guard :: "bool \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b"
  case_nil :: "'a \<Rightarrow> 'b"
  case_cons :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
  case_elem :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
  case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b"

declare [[coercion_args case_guard - + -]]
declare [[coercion_args case_cons - -]]
declare [[coercion_args case_abs -]]
declare [[coercion_args case_elem - +]]

ML_file \<open>Tools/Ctr_Sugar/case_translation.ML\<close>

lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y"
  by (erule iffI) (erule contrapos_pn)

lemma iff_contradict:
  "\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R"
  "\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R"
  by blast+

ML_file \<open>Tools/Ctr_Sugar/ctr_sugar_util.ML\<close>
ML_file \<open>Tools/Ctr_Sugar/ctr_sugar_tactics.ML\<close>
ML_file \<open>Tools/Ctr_Sugar/ctr_sugar_code.ML\<close>
ML_file \<open>Tools/Ctr_Sugar/ctr_sugar.ML\<close>

text \<open>Coinduction method that avoids some boilerplate compared with coinduct.\<close>

ML_file \<open>Tools/coinduction.ML\<close>

end