author | haftmann |
Fri, 20 Oct 2017 20:57:55 +0200 | |
changeset 66888 | 930abfdf8727 |
parent 64430 | 1d85ac286c72 |
child 69605 | a96320074298 |
permissions | -rw-r--r-- |
54397 | 1 |
(* Title: HOL/Ctr_Sugar.thy |
49486 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
55160 | 3 |
Author: Dmitriy Traytel, TU Muenchen |
54397 | 4 |
Copyright 2012, 2013 |
49075 | 5 |
|
51797 | 6 |
Wrapping existing freely generated type's constructors. |
49075 | 7 |
*) |
8 |
||
60758 | 9 |
section \<open>Wrapping Existing Freely Generated Type's Constructors\<close> |
49075 | 10 |
|
54006 | 11 |
theory Ctr_Sugar |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
12 |
imports HOL |
49075 | 13 |
keywords |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
14 |
"print_case_translations" :: diag and |
55468
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
blanchet
parents:
55160
diff
changeset
|
15 |
"free_constructors" :: thy_goal |
49075 | 16 |
begin |
17 |
||
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
18 |
consts |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
19 |
case_guard :: "bool \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b" |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
20 |
case_nil :: "'a \<Rightarrow> 'b" |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
21 |
case_cons :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
22 |
case_elem :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b" |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
23 |
case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b" |
57631
959caab43a3d
use the noted theorem whenever possible, also in 'BNF_Def'
blanchet
parents:
55971
diff
changeset
|
24 |
|
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
25 |
declare [[coercion_args case_guard - + -]] |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
26 |
declare [[coercion_args case_cons - -]] |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
27 |
declare [[coercion_args case_abs -]] |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
28 |
declare [[coercion_args case_elem - +]] |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
29 |
|
54701 | 30 |
ML_file "Tools/Ctr_Sugar/case_translation.ML" |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
31 |
|
49312 | 32 |
lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y" |
58188 | 33 |
by (erule iffI) (erule contrapos_pn) |
49312 | 34 |
|
49486 | 35 |
lemma iff_contradict: |
58188 | 36 |
"\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R" |
37 |
"\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R" |
|
38 |
by blast+ |
|
49486 | 39 |
|
54701 | 40 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_util.ML" |
41 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_tactics.ML" |
|
42 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_code.ML" |
|
43 |
ML_file "Tools/Ctr_Sugar/ctr_sugar.ML" |
|
49309
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
49286
diff
changeset
|
44 |
|
64430
1d85ac286c72
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
blanchet
parents:
63655
diff
changeset
|
45 |
text \<open>Coinduction method that avoids some boilerplate compared with coinduct.\<close> |
63655 | 46 |
|
47 |
ML_file "Tools/coinduction.ML" |
|
48 |
||
49075 | 49 |
end |