| 41828 |      1 | (*  Title:      HOL/ex/Case_Product.thy
 | 
|  |      2 |     Author:     Lars Noschinski
 | 
|  |      3 |     Copyright   2011 TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Examples for the 'case_product' attribute *}
 | 
|  |      7 | 
 | 
|  |      8 | theory Case_Product
 | 
|  |      9 | imports Main
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | text {*
 | 
|  |     13 |   The {@attribute case_product} attribute combines multiple case distinction
 | 
|  |     14 |   lemmas into a single case distinction lemma by building the product of all
 | 
|  |     15 |   these case distinctions.
 | 
|  |     16 | *}
 | 
|  |     17 | 
 | 
|  |     18 | lemmas nat_list_exhaust = nat.exhaust[case_product list.exhaust]
 | 
|  |     19 | 
 | 
|  |     20 | text {*
 | 
|  |     21 |   The attribute honors preconditions
 | 
|  |     22 | *}
 | 
|  |     23 | 
 | 
|  |     24 | lemmas trancl_acc_cases= trancl.cases[case_product acc.cases]
 | 
|  |     25 | 
 | 
|  |     26 | text {*
 | 
|  |     27 |   Also, case names are generated based on the old names
 | 
|  |     28 | *}
 | 
|  |     29 | 
 | 
|  |     30 | end
 |