author | haftmann |
Fri, 27 Dec 2013 14:35:14 +0100 | |
changeset 54868 | bab6cade3cc5 |
parent 54701 | 4ed7454aebde |
child 55160 | 2d69438b1b0c |
permissions | -rw-r--r-- |
54397 | 1 |
(* Title: HOL/Ctr_Sugar.thy |
49486 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
54397 | 3 |
Copyright 2012, 2013 |
49075 | 4 |
|
51797 | 5 |
Wrapping existing freely generated type's constructors. |
49075 | 6 |
*) |
7 |
||
51797 | 8 |
header {* Wrapping Existing Freely Generated Type's Constructors *} |
49075 | 9 |
|
54006 | 10 |
theory Ctr_Sugar |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
11 |
imports HOL |
49075 | 12 |
keywords |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
13 |
"print_case_translations" :: diag and |
54626 | 14 |
"wrap_free_constructors" :: thy_goal |
49075 | 15 |
begin |
16 |
||
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
17 |
consts |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
18 |
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
|
19 |
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
|
20 |
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
|
21 |
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
|
22 |
case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b" |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
23 |
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
|
24 |
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
|
25 |
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
|
26 |
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
|
27 |
|
54701 | 28 |
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
|
29 |
setup Case_Translation.setup |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
54397
diff
changeset
|
30 |
|
49312 | 31 |
lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y" |
32 |
by (erule iffI) (erule contrapos_pn) |
|
33 |
||
49486 | 34 |
lemma iff_contradict: |
35 |
"\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R" |
|
36 |
"\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R" |
|
37 |
by blast+ |
|
38 |
||
54701 | 39 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_util.ML" |
40 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_tactics.ML" |
|
41 |
ML_file "Tools/Ctr_Sugar/ctr_sugar_code.ML" |
|
42 |
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
|
43 |
|
49075 | 44 |
end |